author | berghofe |
Fri, 28 Apr 2006 17:56:20 +0200 | |
changeset 19499 | 1a082c1257d7 |
parent 17429 | e8d6ed3aacfe |
child 19765 | dfe940911617 |
permissions | -rw-r--r-- |
10751 | 1 |
(* Title : HyperPow.thy |
2 |
Author : Jacques D. Fleuriot |
|
3 |
Copyright : 1998 University of Cambridge |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
4 |
Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 |
10751 | 5 |
*) |
6 |
||
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
7 |
header{*Exponentials on the Hyperreals*} |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
8 |
|
15131 | 9 |
theory HyperPow |
15360 | 10 |
imports HyperArith HyperNat |
15131 | 11 |
begin |
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
12 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
13 |
(* consts hpowr :: "[hypreal,nat] => hypreal" *) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
14 |
|
17298 | 15 |
lemma hpowr_0 [simp]: "r ^ 0 = (1::hypreal)" |
16 |
by (rule power_0) |
|
17 |
||
18 |
lemma hpowr_Suc [simp]: "r ^ (Suc n) = (r::hypreal) * (r ^ n)" |
|
19 |
by (rule power_Suc) |
|
10751 | 20 |
|
21 |
consts |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
22 |
"pow" :: "[hypreal,hypnat] => hypreal" (infixr "pow" 80) |
10751 | 23 |
|
24 |
defs |
|
25 |
||
26 |
(* hypernatural powers of hyperreals *) |
|
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
27 |
hyperpow_def [transfer_unfold]: |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
28 |
"(R::hypreal) pow (N::hypnat) == ( *f2* op ^) R N" |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
29 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
30 |
lemma hrealpow_two: "(r::hypreal) ^ Suc (Suc 0) = r * r" |
14443
75910c7557c5
generic theorems about exponentials; general tidying up
paulson
parents:
14435
diff
changeset
|
31 |
by simp |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
32 |
|
15003 | 33 |
lemma hrealpow_two_le [simp]: "(0::hypreal) \<le> r ^ Suc (Suc 0)" |
14371
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents:
14348
diff
changeset
|
34 |
by (auto simp add: zero_le_mult_iff) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
35 |
|
15003 | 36 |
lemma hrealpow_two_le_add_order [simp]: |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
37 |
"(0::hypreal) \<le> u ^ Suc (Suc 0) + v ^ Suc (Suc 0)" |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
38 |
by (simp only: hrealpow_two_le add_nonneg_nonneg) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
39 |
|
15003 | 40 |
lemma hrealpow_two_le_add_order2 [simp]: |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
41 |
"(0::hypreal) \<le> u ^ Suc (Suc 0) + v ^ Suc (Suc 0) + w ^ Suc (Suc 0)" |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
42 |
by (simp only: hrealpow_two_le add_nonneg_nonneg) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
43 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
44 |
lemma hypreal_add_nonneg_eq_0_iff: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
45 |
"[| 0 \<le> x; 0 \<le> y |] ==> (x+y = 0) = (x = 0 & y = (0::hypreal))" |
15003 | 46 |
by arith |
47 |
||
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
48 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
49 |
text{*FIXME: DELETE THESE*} |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
50 |
lemma hypreal_three_squares_add_zero_iff: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
51 |
"(x*x + y*y + z*z = 0) = (x = 0 & y = 0 & z = (0::hypreal))" |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
52 |
apply (simp only: zero_le_square add_nonneg_nonneg hypreal_add_nonneg_eq_0_iff, auto) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
53 |
done |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
54 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
55 |
lemma hrealpow_three_squares_add_zero_iff [simp]: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
56 |
"(x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + z ^ Suc (Suc 0) = (0::hypreal)) = |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
57 |
(x = 0 & y = 0 & z = 0)" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
58 |
by (simp only: hypreal_three_squares_add_zero_iff hrealpow_two) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
59 |
|
16924 | 60 |
(*FIXME: This and RealPow.abs_realpow_two should be replaced by an abstract |
61 |
result proved in Ring_and_Field*) |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
62 |
lemma hrabs_hrealpow_two [simp]: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
63 |
"abs(x ^ Suc (Suc 0)) = (x::hypreal) ^ Suc (Suc 0)" |
16924 | 64 |
by (simp add: abs_mult) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
65 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
66 |
lemma two_hrealpow_ge_one [simp]: "(1::hypreal) \<le> 2 ^ n" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
67 |
by (insert power_increasing [of 0 n "2::hypreal"], simp) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
68 |
|
15003 | 69 |
lemma two_hrealpow_gt [simp]: "hypreal_of_nat n < 2 ^ n" |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
70 |
apply (induct_tac "n") |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
71 |
apply (auto simp add: hypreal_of_nat_Suc left_distrib) |
14371
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents:
14348
diff
changeset
|
72 |
apply (cut_tac n = n in two_hrealpow_ge_one, arith) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
73 |
done |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
74 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
75 |
lemma hrealpow: |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
76 |
"star_n X ^ m = star_n (%n. (X n::real) ^ m)" |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
77 |
apply (induct_tac "m") |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
78 |
apply (auto simp add: star_n_one_num star_n_mult power_0) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
79 |
done |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
80 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
81 |
lemma hrealpow_sum_square_expand: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
82 |
"(x + (y::hypreal)) ^ Suc (Suc 0) = |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
83 |
x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + (hypreal_of_nat (Suc (Suc 0)))*x*y" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
84 |
by (simp add: right_distrib left_distrib hypreal_of_nat_Suc) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
85 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
86 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
87 |
subsection{*Literal Arithmetic Involving Powers and Type @{typ hypreal}*} |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
88 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
89 |
lemma power_hypreal_of_real_number_of: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
90 |
"(number_of v :: hypreal) ^ n = hypreal_of_real ((number_of v) ^ n)" |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
91 |
by simp |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
92 |
(* why is this a simp rule? - BH *) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
93 |
declare power_hypreal_of_real_number_of [of _ "number_of w", standard, simp] |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
94 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
95 |
lemma hrealpow_HFinite: "x \<in> HFinite ==> x ^ n \<in> HFinite" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
96 |
apply (induct_tac "n") |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
97 |
apply (auto intro: HFinite_mult) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
98 |
done |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
99 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
100 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
101 |
subsection{*Powers with Hypernatural Exponents*} |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
102 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
103 |
lemma hyperpow: "star_n X pow star_n Y = star_n (%n. X n ^ Y n)" |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
104 |
by (simp add: hyperpow_def starfun2_star_n) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
105 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
106 |
lemma hyperpow_zero [simp]: "!!n. (0::hypreal) pow (n + (1::hypnat)) = 0" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
107 |
by (transfer, simp) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
108 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
109 |
lemma hyperpow_not_zero: "!!r n. r \<noteq> (0::hypreal) ==> r pow n \<noteq> 0" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
110 |
by (transfer, simp) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
111 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
112 |
lemma hyperpow_inverse: |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
113 |
"!!r n. r \<noteq> (0::hypreal) ==> inverse(r pow n) = (inverse r) pow n" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
114 |
by (transfer, rule power_inverse) |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
115 |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
116 |
lemma hyperpow_hrabs: "!!r n. abs r pow n = abs (r pow n)" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
117 |
by (transfer, rule power_abs [symmetric]) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
118 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
119 |
lemma hyperpow_add: "!!r n m. r pow (n + m) = (r pow n) * (r pow m)" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
120 |
by (transfer, rule power_add) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
121 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
122 |
lemma hyperpow_one [simp]: "!!r. r pow (1::hypnat) = r" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
123 |
by (transfer, simp) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
124 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
125 |
lemma hyperpow_two: |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
126 |
"!!r. r pow ((1::hypnat) + (1::hypnat)) = r * r" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
127 |
by (transfer, simp) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
128 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
129 |
lemma hyperpow_gt_zero: "!!r n. (0::hypreal) < r ==> 0 < r pow n" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
130 |
by (transfer, rule zero_less_power) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
131 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
132 |
lemma hyperpow_ge_zero: "!!r n. (0::hypreal) \<le> r ==> 0 \<le> r pow n" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
133 |
by (transfer, rule zero_le_power) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
134 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
135 |
lemma hyperpow_le: |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
136 |
"!!x y n. [|(0::hypreal) < x; x \<le> y|] ==> x pow n \<le> y pow n" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
137 |
by (transfer, rule power_mono [OF _ order_less_imp_le]) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
138 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
139 |
lemma hyperpow_eq_one [simp]: "!!n. 1 pow n = (1::hypreal)" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
140 |
by (transfer, simp) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
141 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
142 |
lemma hrabs_hyperpow_minus_one [simp]: "!!n. abs(-1 pow n) = (1::hypreal)" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
143 |
by (transfer, simp) |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
144 |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
145 |
lemma hyperpow_mult: "!!r s n. (r * s) pow n = (r pow n) * (s pow n)" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
146 |
by (transfer, rule power_mult_distrib) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
147 |
|
15003 | 148 |
lemma hyperpow_two_le [simp]: "0 \<le> r pow (1 + 1)" |
14371
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents:
14348
diff
changeset
|
149 |
by (auto simp add: hyperpow_two zero_le_mult_iff) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
150 |
|
14371
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents:
14348
diff
changeset
|
151 |
lemma hrabs_hyperpow_two [simp]: "abs(x pow (1 + 1)) = x pow (1 + 1)" |
15003 | 152 |
by (simp add: abs_if hyperpow_two_le linorder_not_less) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
153 |
|
15003 | 154 |
lemma hyperpow_two_hrabs [simp]: "abs(x) pow (1 + 1) = x pow (1 + 1)" |
155 |
by (simp add: hyperpow_hrabs) |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
156 |
|
15229 | 157 |
text{*The precondition could be weakened to @{term "0\<le>x"}*} |
158 |
lemma hypreal_mult_less_mono: |
|
159 |
"[| u<v; x<y; (0::hypreal) < v; 0 < x |] ==> u*x < v* y" |
|
160 |
by (simp add: Ring_and_Field.mult_strict_mono order_less_imp_le) |
|
161 |
||
15003 | 162 |
lemma hyperpow_two_gt_one: "1 < r ==> 1 < r pow (1 + 1)" |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
163 |
apply (auto simp add: hyperpow_two) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
164 |
apply (rule_tac y = "1*1" in order_le_less_trans) |
14371
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents:
14348
diff
changeset
|
165 |
apply (rule_tac [2] hypreal_mult_less_mono, auto) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
166 |
done |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
167 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
168 |
lemma hyperpow_two_ge_one: |
15003 | 169 |
"1 \<le> r ==> 1 \<le> r pow (1 + 1)" |
170 |
by (auto dest!: order_le_imp_less_or_eq intro: hyperpow_two_gt_one order_less_imp_le) |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
171 |
|
15003 | 172 |
lemma two_hyperpow_ge_one [simp]: "(1::hypreal) \<le> 2 pow n" |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
173 |
apply (rule_tac y = "1 pow n" in order_trans) |
14371
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents:
14348
diff
changeset
|
174 |
apply (rule_tac [2] hyperpow_le, auto) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
175 |
done |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
176 |
|
15003 | 177 |
lemma hyperpow_minus_one2 [simp]: |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
178 |
"!!n. -1 pow ((1 + 1)*n) = (1::hypreal)" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
179 |
by (transfer, simp) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
180 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
181 |
lemma hyperpow_less_le: |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
182 |
"!!r n N. [|(0::hypreal) \<le> r; r \<le> 1; n < N|] ==> r pow N \<le> r pow n" |
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17318
diff
changeset
|
183 |
by (transfer, rule power_decreasing [OF order_less_imp_le]) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
184 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
185 |
lemma hyperpow_SHNat_le: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
186 |
"[| 0 \<le> r; r \<le> (1::hypreal); N \<in> HNatInfinite |] |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
187 |
==> ALL n: Nats. r pow N \<le> r pow n" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
188 |
by (auto intro!: hyperpow_less_le simp add: HNatInfinite_iff) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
189 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
190 |
lemma hyperpow_realpow: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
191 |
"(hypreal_of_real r) pow (hypnat_of_nat n) = hypreal_of_real (r ^ n)" |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
192 |
by (simp add: star_of_def hypnat_of_nat_eq hyperpow) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
193 |
|
15003 | 194 |
lemma hyperpow_SReal [simp]: |
195 |
"(hypreal_of_real r) pow (hypnat_of_nat n) \<in> Reals" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
196 |
by (simp del: star_of_power add: hyperpow_realpow SReal_def) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
197 |
|
15003 | 198 |
|
199 |
lemma hyperpow_zero_HNatInfinite [simp]: |
|
200 |
"N \<in> HNatInfinite ==> (0::hypreal) pow N = 0" |
|
14371
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents:
14348
diff
changeset
|
201 |
by (drule HNatInfinite_is_Suc, auto) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
202 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
203 |
lemma hyperpow_le_le: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
204 |
"[| (0::hypreal) \<le> r; r \<le> 1; n \<le> N |] ==> r pow N \<le> r pow n" |
14371
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents:
14348
diff
changeset
|
205 |
apply (drule order_le_less [of n, THEN iffD1]) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
206 |
apply (auto intro: hyperpow_less_le) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
207 |
done |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
208 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
209 |
lemma hyperpow_Suc_le_self2: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
210 |
"[| (0::hypreal) \<le> r; r < 1 |] ==> r pow (n + (1::hypnat)) \<le> r" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
211 |
apply (drule_tac n = " (1::hypnat) " in hyperpow_le_le) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
212 |
apply auto |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
213 |
done |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
214 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
215 |
lemma lemma_Infinitesimal_hyperpow: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
216 |
"[| x \<in> Infinitesimal; 0 < N |] ==> abs (x pow N) \<le> abs x" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
217 |
apply (unfold Infinitesimal_def) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
218 |
apply (auto intro!: hyperpow_Suc_le_self2 |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
219 |
simp add: hyperpow_hrabs [symmetric] hypnat_gt_zero_iff2 abs_ge_zero) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
220 |
done |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
221 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
222 |
lemma Infinitesimal_hyperpow: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
223 |
"[| x \<in> Infinitesimal; 0 < N |] ==> x pow N \<in> Infinitesimal" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
224 |
apply (rule hrabs_le_Infinitesimal) |
14371
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents:
14348
diff
changeset
|
225 |
apply (rule_tac [2] lemma_Infinitesimal_hyperpow, auto) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
226 |
done |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
227 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
228 |
lemma hrealpow_hyperpow_Infinitesimal_iff: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
229 |
"(x ^ n \<in> Infinitesimal) = (x pow (hypnat_of_nat n) \<in> Infinitesimal)" |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
230 |
apply (cases x) |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
231 |
apply (simp add: hrealpow hyperpow hypnat_of_nat_eq) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
232 |
done |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
233 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
234 |
lemma Infinitesimal_hrealpow: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
235 |
"[| x \<in> Infinitesimal; 0 < n |] ==> x ^ n \<in> Infinitesimal" |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
236 |
by (simp add: hrealpow_hyperpow_Infinitesimal_iff Infinitesimal_hyperpow) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
237 |
|
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
238 |
ML |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
239 |
{* |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
240 |
val hrealpow_two = thm "hrealpow_two"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
241 |
val hrealpow_two_le = thm "hrealpow_two_le"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
242 |
val hrealpow_two_le_add_order = thm "hrealpow_two_le_add_order"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
243 |
val hrealpow_two_le_add_order2 = thm "hrealpow_two_le_add_order2"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
244 |
val hypreal_add_nonneg_eq_0_iff = thm "hypreal_add_nonneg_eq_0_iff"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
245 |
val hypreal_three_squares_add_zero_iff = thm "hypreal_three_squares_add_zero_iff"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
246 |
val hrealpow_three_squares_add_zero_iff = thm "hrealpow_three_squares_add_zero_iff"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
247 |
val hrabs_hrealpow_two = thm "hrabs_hrealpow_two"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
248 |
val two_hrealpow_ge_one = thm "two_hrealpow_ge_one"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
249 |
val two_hrealpow_gt = thm "two_hrealpow_gt"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
250 |
val hrealpow = thm "hrealpow"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
251 |
val hrealpow_sum_square_expand = thm "hrealpow_sum_square_expand"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
252 |
val power_hypreal_of_real_number_of = thm "power_hypreal_of_real_number_of"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
253 |
val hrealpow_HFinite = thm "hrealpow_HFinite"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
254 |
val hyperpow = thm "hyperpow"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
255 |
val hyperpow_zero = thm "hyperpow_zero"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
256 |
val hyperpow_not_zero = thm "hyperpow_not_zero"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
257 |
val hyperpow_inverse = thm "hyperpow_inverse"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
258 |
val hyperpow_hrabs = thm "hyperpow_hrabs"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
259 |
val hyperpow_add = thm "hyperpow_add"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
260 |
val hyperpow_one = thm "hyperpow_one"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
261 |
val hyperpow_two = thm "hyperpow_two"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
262 |
val hyperpow_gt_zero = thm "hyperpow_gt_zero"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
263 |
val hyperpow_ge_zero = thm "hyperpow_ge_zero"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
264 |
val hyperpow_le = thm "hyperpow_le"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
265 |
val hyperpow_eq_one = thm "hyperpow_eq_one"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
266 |
val hrabs_hyperpow_minus_one = thm "hrabs_hyperpow_minus_one"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
267 |
val hyperpow_mult = thm "hyperpow_mult"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
268 |
val hyperpow_two_le = thm "hyperpow_two_le"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
269 |
val hrabs_hyperpow_two = thm "hrabs_hyperpow_two"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
270 |
val hyperpow_two_hrabs = thm "hyperpow_two_hrabs"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
271 |
val hyperpow_two_gt_one = thm "hyperpow_two_gt_one"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
272 |
val hyperpow_two_ge_one = thm "hyperpow_two_ge_one"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
273 |
val two_hyperpow_ge_one = thm "two_hyperpow_ge_one"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
274 |
val hyperpow_minus_one2 = thm "hyperpow_minus_one2"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
275 |
val hyperpow_less_le = thm "hyperpow_less_le"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
276 |
val hyperpow_SHNat_le = thm "hyperpow_SHNat_le"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
277 |
val hyperpow_realpow = thm "hyperpow_realpow"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
278 |
val hyperpow_SReal = thm "hyperpow_SReal"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
279 |
val hyperpow_zero_HNatInfinite = thm "hyperpow_zero_HNatInfinite"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
280 |
val hyperpow_le_le = thm "hyperpow_le_le"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
281 |
val hyperpow_Suc_le_self2 = thm "hyperpow_Suc_le_self2"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
282 |
val lemma_Infinitesimal_hyperpow = thm "lemma_Infinitesimal_hyperpow"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
283 |
val Infinitesimal_hyperpow = thm "Infinitesimal_hyperpow"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
284 |
val hrealpow_hyperpow_Infinitesimal_iff = thm "hrealpow_hyperpow_Infinitesimal_iff"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
285 |
val Infinitesimal_hrealpow = thm "Infinitesimal_hrealpow"; |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
286 |
*} |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
11713
diff
changeset
|
287 |
|
10751 | 288 |
end |