src/HOL/Hyperreal/HyperArith.thy
author paulson
Tue, 10 Feb 2004 12:02:11 +0100
changeset 14378 69c4d5997669
parent 14371 c78c7da09519
child 14387 e96d5c42c4b0
permissions -rw-r--r--
generic of_nat and of_int functions, and generalization of iszero and neg
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
     1
(*  Title:      HOL/HyperArith.thy
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
     2
    ID:         $Id$
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
     4
    Copyright   1999  University of Cambridge
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
     5
*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
     6
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
     7
header{*Binary arithmetic and Simplification for the Hyperreals*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
     8
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
     9
theory HyperArith = HyperDef
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    10
files ("hypreal_arith.ML"):
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    11
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    12
subsection{*Binary Arithmetic for the Hyperreals*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    13
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    14
instance hypreal :: number ..
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    15
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    16
defs (overloaded)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    17
  hypreal_number_of_def:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    18
    "number_of v == hypreal_of_real (number_of v)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    19
     (*::bin=>hypreal               ::bin=>real*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    20
     --{*This case is reduced to that for the reals.*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    21
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    22
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    23
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    24
subsubsection{*Embedding the Reals into the Hyperreals*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    25
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    26
lemma hypreal_number_of [simp]: "hypreal_of_real (number_of w) = number_of w"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    27
by (simp add: hypreal_number_of_def)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    28
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    29
lemma hypreal_numeral_0_eq_0: "Numeral0 = (0::hypreal)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    30
by (simp add: hypreal_number_of_def)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    31
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    32
lemma hypreal_numeral_1_eq_1: "Numeral1 = (1::hypreal)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    33
by (simp add: hypreal_number_of_def)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    34
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    35
subsubsection{*Addition*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    36
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    37
lemma add_hypreal_number_of [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    38
     "(number_of v :: hypreal) + number_of v' = number_of (bin_add v v')"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    39
by (simp only: hypreal_number_of_def hypreal_of_real_add [symmetric]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    40
               add_real_number_of)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    41
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    42
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    43
subsubsection{*Subtraction*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    44
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    45
lemma minus_hypreal_number_of [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    46
     "- (number_of w :: hypreal) = number_of (bin_minus w)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    47
by (simp only: hypreal_number_of_def minus_real_number_of
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    48
               hypreal_of_real_minus [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    49
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    50
lemma diff_hypreal_number_of [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    51
     "(number_of v :: hypreal) - number_of w =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    52
      number_of (bin_add v (bin_minus w))"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    53
by (unfold hypreal_number_of_def hypreal_diff_def, simp)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    54
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    55
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    56
subsubsection{*Multiplication*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    57
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    58
lemma mult_hypreal_number_of [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    59
     "(number_of v :: hypreal) * number_of v' = number_of (bin_mult v v')"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    60
by (simp only: hypreal_number_of_def hypreal_of_real_mult [symmetric] mult_real_number_of)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    61
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    62
text{*Lemmas for specialist use, NOT as default simprules*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    63
lemma hypreal_mult_2: "2 * z = (z+z::hypreal)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    64
proof -
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    65
  have eq: "(2::hypreal) = 1 + 1"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    66
    by (simp add: hypreal_numeral_1_eq_1 [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    67
  thus ?thesis by (simp add: eq left_distrib)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    68
qed
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    69
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    70
lemma hypreal_mult_2_right: "z * 2 = (z+z::hypreal)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    71
by (subst hypreal_mult_commute, rule hypreal_mult_2)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    72
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    73
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    74
subsubsection{*Comparisons*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    75
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    76
(** Equals (=) **)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    77
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    78
lemma eq_hypreal_number_of [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    79
     "((number_of v :: hypreal) = number_of v') =
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
    80
      iszero (number_of (bin_add v (bin_minus v')) :: int)"
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    81
apply (simp only: hypreal_number_of_def hypreal_of_real_eq_iff eq_real_number_of)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    82
done
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    83
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    84
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    85
(** Less-than (<) **)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    86
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    87
(*"neg" is used in rewrite rules for binary comparisons*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    88
lemma less_hypreal_number_of [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    89
     "((number_of v :: hypreal) < number_of v') =
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
    90
      neg (number_of (bin_add v (bin_minus v')) :: int)"
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    91
by (simp only: hypreal_number_of_def hypreal_of_real_less_iff less_real_number_of)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    92
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    93
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    94
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    95
text{*New versions of existing theorems involving 0, 1*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    96
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    97
lemma hypreal_minus_1_eq_m1 [simp]: "- 1 = (-1::hypreal)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    98
by (simp add: hypreal_numeral_1_eq_1 [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    99
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   100
lemma hypreal_mult_minus1 [simp]: "-1 * z = -(z::hypreal)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   101
proof -
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   102
  have  "-1 * z = (- 1) * z" by (simp add: hypreal_minus_1_eq_m1)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   103
  also have "... = - (1 * z)" by (simp only: minus_mult_left)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   104
  also have "... = -z" by simp
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   105
  finally show ?thesis .
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   106
qed
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   107
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   108
lemma hypreal_mult_minus1_right [simp]: "(z::hypreal) * -1 = -z"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   109
by (subst hypreal_mult_commute, rule hypreal_mult_minus1)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   110
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   111
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   112
subsection{*Simplification of Arithmetic when Nested to the Right*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   113
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   114
lemma hypreal_add_number_of_left [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   115
     "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::hypreal)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   116
by (simp add: add_assoc [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   117
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   118
lemma hypreal_mult_number_of_left [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   119
     "number_of v *(number_of w * z) = (number_of(bin_mult v w) * z::hypreal)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   120
by (simp add: hypreal_mult_assoc [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   121
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   122
lemma hypreal_add_number_of_diff1:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   123
    "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::hypreal)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   124
by (simp add: hypreal_diff_def hypreal_add_number_of_left)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   125
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   126
lemma hypreal_add_number_of_diff2 [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   127
     "number_of v + (c - number_of w) =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   128
      number_of (bin_add v (bin_minus w)) + (c::hypreal)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   129
apply (subst diff_hypreal_number_of [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   130
apply (simp only: hypreal_diff_def add_ac)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   131
done
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   132
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   133
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   134
declare hypreal_numeral_0_eq_0 [simp] hypreal_numeral_1_eq_1 [simp]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   135
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   136
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   137
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   138
use "hypreal_arith.ML"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   139
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   140
setup hypreal_arith_setup
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   141
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 14369
diff changeset
   142
lemma hypreal_le_add_order: "[| 0 \<le> x; 0 \<le> y |] ==> (0::hypreal) \<le> x + y"
b0064703967b simplifications in the hyperreals
paulson
parents: 14369
diff changeset
   143
by arith
b0064703967b simplifications in the hyperreals
paulson
parents: 14369
diff changeset
   144
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14309
diff changeset
   145
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   146
subsubsection{*Division By @{term 1} and @{term "-1"}*}
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   147
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   148
lemma hypreal_divide_minus1 [simp]: "x/-1 = -(x::hypreal)"
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   149
by simp
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   150
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   151
lemma hypreal_minus1_divide [simp]: "-1/(x::hypreal) = - (1/x)"
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   152
by (simp add: hypreal_divide_def hypreal_minus_inverse)
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   153
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   154
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   155
subsection{*The Function @{term hypreal_of_real}*}
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   156
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   157
lemma number_of_less_hypreal_of_real_iff [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   158
     "(number_of w < hypreal_of_real z) = (number_of w < z)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   159
apply (subst hypreal_of_real_less_iff [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   160
apply (simp (no_asm))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   161
done
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   162
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   163
lemma number_of_le_hypreal_of_real_iff [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   164
     "(number_of w <= hypreal_of_real z) = (number_of w <= z)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   165
apply (subst hypreal_of_real_le_iff [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   166
apply (simp (no_asm))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   167
done
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   168
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   169
lemma hypreal_of_real_eq_number_of_iff [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   170
     "(hypreal_of_real z = number_of w) = (z = number_of w)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   171
apply (subst hypreal_of_real_eq_iff [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   172
apply (simp (no_asm))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   173
done
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   174
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   175
lemma hypreal_of_real_less_number_of_iff [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   176
     "(hypreal_of_real z < number_of w) = (z < number_of w)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   177
apply (subst hypreal_of_real_less_iff [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   178
apply (simp (no_asm))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   179
done
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   180
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   181
lemma hypreal_of_real_le_number_of_iff [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   182
     "(hypreal_of_real z <= number_of w) = (z <= number_of w)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   183
apply (subst hypreal_of_real_le_iff [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   184
apply (simp (no_asm))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   185
done
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   186
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   187
subsection{*Absolute Value Function for the Hyperreals*}
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   188
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   189
lemma hrabs_number_of [simp]:
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   190
     "abs (number_of v :: hypreal) =
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   191
        (if neg (number_of v :: int) then number_of (bin_minus v)
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   192
         else number_of v)"
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   193
by (simp add: hrabs_def)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   194
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   195
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   196
declare abs_mult [simp]
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   197
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   198
lemma hrabs_add_less: "[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)"
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   199
apply (unfold hrabs_def)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   200
apply (simp split add: split_if_asm)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   201
done
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   202
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   203
lemma hrabs_less_gt_zero: "abs x < r ==> (0::hypreal) < r"
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   204
by (blast intro!: order_le_less_trans abs_ge_zero)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   205
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   206
lemma hrabs_disj: "abs x = (x::hypreal) | abs x = -x"
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   207
by (simp add: hrabs_def)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   208
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   209
(* Needed in Geom.ML *)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   210
lemma hrabs_add_lemma_disj: "(y::hypreal) + - x + (y + - z) = abs (x + - z) ==> y = z | x = y"
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   211
by (simp add: hrabs_def split add: split_if_asm)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   212
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   213
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   214
(*----------------------------------------------------------
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   215
    Relating hrabs to abs through embedding of IR into IR*
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   216
 ----------------------------------------------------------*)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   217
lemma hypreal_of_real_hrabs:
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   218
    "abs (hypreal_of_real r) = hypreal_of_real (abs r)"
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   219
apply (unfold hypreal_of_real_def)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   220
apply (auto simp add: hypreal_hrabs)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   221
done
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   222
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   223
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   224
subsection{*Embedding the Naturals into the Hyperreals*}
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   225
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   226
constdefs
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   227
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   228
  hypreal_of_nat   :: "nat => hypreal"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   229
   "hypreal_of_nat m  == of_nat m"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   230
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   231
lemma SNat_eq: "Nats = {n. \<exists>N. n = hypreal_of_nat N}"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   232
by (force simp add: hypreal_of_nat_def Nats_def) 
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   233
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   234
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   235
lemma hypreal_of_nat_add [simp]:
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   236
     "hypreal_of_nat (m + n) = hypreal_of_nat m + hypreal_of_nat n"
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   237
by (simp add: hypreal_of_nat_def)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   238
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   239
lemma hypreal_of_nat_mult: "hypreal_of_nat (m * n) = hypreal_of_nat m * hypreal_of_nat n"
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   240
by (simp add: hypreal_of_nat_def)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   241
declare hypreal_of_nat_mult [simp]
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   242
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   243
lemma hypreal_of_nat_less_iff:
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   244
      "(n < m) = (hypreal_of_nat n < hypreal_of_nat m)"
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   245
apply (simp add: hypreal_of_nat_def)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   246
done
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   247
declare hypreal_of_nat_less_iff [symmetric, simp]
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   248
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   249
(*------------------------------------------------------------*)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   250
(* naturals embedded in hyperreals                            *)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   251
(* is a hyperreal c.f. NS extension                           *)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   252
(*------------------------------------------------------------*)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   253
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   254
lemma hypreal_of_nat_eq:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   255
     "hypreal_of_nat (n::nat) = hypreal_of_real (real n)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   256
apply (induct n) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   257
apply (simp_all add: hypreal_of_nat_def real_of_nat_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   258
done
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   259
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   260
lemma hypreal_of_nat:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   261
     "hypreal_of_nat m = Abs_hypreal(hyprel``{%n. real m})"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   262
apply (induct m) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   263
apply (simp_all add: hypreal_of_nat_def real_of_nat_def hypreal_zero_def 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   264
                     hypreal_one_def hypreal_add)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   265
done
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   266
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   267
lemma hypreal_of_nat_Suc:
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   268
     "hypreal_of_nat (Suc n) = hypreal_of_nat n + (1::hypreal)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   269
by (simp add: hypreal_of_nat_def)
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   270
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   271
(*"neg" is used in rewrite rules for binary comparisons*)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   272
lemma hypreal_of_nat_number_of [simp]:
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   273
     "hypreal_of_nat (number_of v :: nat) =
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   274
         (if neg (number_of v :: int) then 0
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   275
          else (number_of v :: hypreal))"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   276
by (simp add: hypreal_of_nat_eq)
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   277
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   278
lemma hypreal_of_nat_zero [simp]: "hypreal_of_nat 0 = 0"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   279
by (simp add: hypreal_of_nat_def) 
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   280
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   281
lemma hypreal_of_nat_one [simp]: "hypreal_of_nat 1 = 1"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   282
by (simp add: hypreal_of_nat_def) 
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   283
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   284
lemma hypreal_of_nat_le_iff [simp]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   285
     "(hypreal_of_nat n \<le> hypreal_of_nat m) = (n \<le> m)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   286
by (simp add: hypreal_of_nat_def) 
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   287
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   288
lemma hypreal_of_nat_ge_zero [simp]: "0 \<le> hypreal_of_nat n"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   289
by (simp add: hypreal_of_nat_def) 
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   290
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   291
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   292
(*
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   293
FIXME: we should declare this, as for type int, but many proofs would break.
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   294
It replaces x+-y by x-y.
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   295
Addsimps [symmetric hypreal_diff_def]
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   296
*)
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   297
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   298
ML
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   299
{*
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 14369
diff changeset
   300
val hypreal_le_add_order = thm"hypreal_le_add_order";
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   301
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   302
val hypreal_of_nat_def = thm"hypreal_of_nat_def";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   303
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   304
val hrabs_number_of = thm "hrabs_number_of";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   305
val hrabs_add_less = thm "hrabs_add_less";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   306
val hrabs_less_gt_zero = thm "hrabs_less_gt_zero";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   307
val hrabs_disj = thm "hrabs_disj";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   308
val hrabs_add_lemma_disj = thm "hrabs_add_lemma_disj";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   309
val hypreal_of_real_hrabs = thm "hypreal_of_real_hrabs";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   310
val hypreal_of_nat_add = thm "hypreal_of_nat_add";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   311
val hypreal_of_nat_mult = thm "hypreal_of_nat_mult";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   312
val hypreal_of_nat_less_iff = thm "hypreal_of_nat_less_iff";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   313
val hypreal_of_nat_Suc = thm "hypreal_of_nat_Suc";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   314
val hypreal_of_nat_number_of = thm "hypreal_of_nat_number_of";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   315
val hypreal_of_nat_zero = thm "hypreal_of_nat_zero";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   316
val hypreal_of_nat_one = thm "hypreal_of_nat_one";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   317
val hypreal_of_nat_le_iff = thm"hypreal_of_nat_le_iff";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   318
val hypreal_of_nat_ge_zero = thm"hypreal_of_nat_ge_zero";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   319
val hypreal_of_nat = thm"hypreal_of_nat";
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   320
*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   321
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   322
end