src/HOL/Hyperreal/HyperArith.thy
author obua
Sun, 09 May 2004 23:04:36 +0200
changeset 14722 8e739a6eaf11
parent 14387 e96d5c42c4b0
child 15003 6145dd7538d7
permissions -rw-r--r--
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
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
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    12
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    13
subsection{*Numerals and Arithmetic*}
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    14
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    15
instance hypreal :: number ..
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    16
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    17
primrec (*the type constraint is essential!*)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    18
  number_of_Pls: "number_of bin.Pls = 0"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    19
  number_of_Min: "number_of bin.Min = - (1::hypreal)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    20
  number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) +
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    21
	                               (number_of w) + (number_of w)"
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    22
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    23
declare number_of_Pls [simp del]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    24
        number_of_Min [simp del]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    25
        number_of_BIT [simp del]
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    26
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    27
instance hypreal :: number_ring
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    28
proof
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    29
  show "Numeral0 = (0::hypreal)" by (rule number_of_Pls)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    30
  show "-1 = - (1::hypreal)" by (rule number_of_Min)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    31
  fix w :: bin and x :: bool
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    32
  show "(number_of (w BIT x) :: hypreal) =
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    33
        (if x then 1 else 0) + number_of w + number_of w"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    34
    by (rule number_of_BIT)
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    35
qed
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
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    38
text{*Collapse applications of @{term hypreal_of_real} to @{term number_of}*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    39
lemma hypreal_number_of [simp]: "hypreal_of_real (number_of w) = number_of w"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    40
apply (induct w) 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    41
apply (simp_all only: number_of hypreal_of_real_add hypreal_of_real_minus, simp_all) 
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    42
done
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    43
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
use "hypreal_arith.ML"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    46
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    47
setup hypreal_arith_setup
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    48
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 14369
diff changeset
    49
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
    50
by arith
b0064703967b simplifications in the hyperreals
paulson
parents: 14369
diff changeset
    51
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14309
diff changeset
    52
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
    53
subsection{*The Function @{term hypreal_of_real}*}
14369
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
lemma number_of_less_hypreal_of_real_iff [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    56
     "(number_of w < hypreal_of_real z) = (number_of w < z)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    57
apply (subst hypreal_of_real_less_iff [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    58
apply (simp (no_asm))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    59
done
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    60
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    61
lemma number_of_le_hypreal_of_real_iff [simp]:
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    62
     "(number_of w \<le> hypreal_of_real z) = (number_of w \<le> z)"
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    63
apply (subst hypreal_of_real_le_iff [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    64
apply (simp (no_asm))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    65
done
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    66
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    67
lemma hypreal_of_real_eq_number_of_iff [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    68
     "(hypreal_of_real z = number_of w) = (z = number_of w)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    69
apply (subst hypreal_of_real_eq_iff [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    70
apply (simp (no_asm))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    71
done
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
lemma hypreal_of_real_less_number_of_iff [simp]:
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    74
     "(hypreal_of_real z < number_of w) = (z < number_of w)"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    75
apply (subst hypreal_of_real_less_iff [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    76
apply (simp (no_asm))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    77
done
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    78
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    79
lemma hypreal_of_real_le_number_of_iff [simp]:
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    80
     "(hypreal_of_real z \<le> number_of w) = (z \<le> number_of w)"
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    81
apply (subst hypreal_of_real_le_iff [symmetric])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    82
apply (simp (no_asm))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    83
done
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    84
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
    85
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
    86
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
    87
declare abs_mult [simp]
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
    88
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
    89
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
    90
apply (unfold hrabs_def)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
    91
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
    92
done
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
    93
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    94
text{*used once in NSA*}
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
    95
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
    96
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
    97
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
    98
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
    99
by (simp add: hrabs_def)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   100
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   101
(* Needed in Geom.ML *)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   102
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
   103
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
   104
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   105
lemma hypreal_of_real_hrabs:
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   106
    "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
   107
apply (unfold hypreal_of_real_def)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   108
apply (auto simp add: hypreal_hrabs)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   109
done
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   110
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   111
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   112
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
   113
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   114
constdefs
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   115
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   116
  hypreal_of_nat   :: "nat => hypreal"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   117
   "hypreal_of_nat m  == of_nat m"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   118
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   119
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
   120
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
   121
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   122
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   123
lemma hypreal_of_nat_add [simp]:
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   124
     "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
   125
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
   126
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   127
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
   128
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
   129
declare hypreal_of_nat_mult [simp]
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   130
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   131
lemma hypreal_of_nat_less_iff:
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   132
      "(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
   133
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
   134
done
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   135
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
   136
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   137
(*------------------------------------------------------------*)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   138
(* naturals embedded in hyperreals                            *)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   139
(* 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
   140
(*------------------------------------------------------------*)
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   141
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   142
lemma hypreal_of_nat_eq:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   143
     "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
   144
apply (induct n) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   145
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
   146
done
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   147
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   148
lemma hypreal_of_nat:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   149
     "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
   150
apply (induct m) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   151
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
   152
                     hypreal_one_def hypreal_add)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   153
done
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   154
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   155
lemma hypreal_of_nat_Suc:
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   156
     "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
   157
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
   158
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   159
(*"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
   160
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
   161
     "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
   162
         (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
   163
          else (number_of v :: hypreal))"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   164
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
   165
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   166
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
   167
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
   168
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   169
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
   170
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
   171
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   172
lemma hypreal_of_nat_le_iff [simp]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   173
     "(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
   174
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
   175
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   176
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
   177
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
   178
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   179
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   180
(*
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   181
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
   182
It replaces x+-y by x-y.
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   183
Addsimps [symmetric hypreal_diff_def]
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   184
*)
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 10751
diff changeset
   185
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   186
ML
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   187
{*
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 14369
diff changeset
   188
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
   189
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   190
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
   191
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   192
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
   193
val hrabs_disj = thm "hrabs_disj";
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   194
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
   195
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
   196
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
   197
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
   198
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
   199
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
   200
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
   201
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
   202
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
   203
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
   204
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
   205
val hypreal_of_nat = thm"hypreal_of_nat";
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   206
*}
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   207
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   208
end