| author | haftmann | 
| Fri, 14 Mar 2008 08:52:51 +0100 | |
| changeset 26266 | 35ae83ca190a | 
| parent 21865 | 55cc354fd2d9 | 
| child 26806 | 40b411ec05aa | 
| permissions | -rw-r--r-- | 
| 10751 | 1 | (* Title : NatStar.thy | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 1998 University of Cambridge | |
| 14415 | 4 | |
| 5 | Converted to Isar and polished by lcp | |
| 6 | *) | |
| 7 | ||
| 8 | header{*Star-transforms for the Hypernaturals*}
 | |
| 10751 | 9 | |
| 15131 | 10 | theory NatStar | 
| 21865 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21864diff
changeset | 11 | imports Star | 
| 15131 | 12 | begin | 
| 14415 | 13 | |
| 17443 | 14 | lemma star_n_eq_starfun_whn: "star_n X = ( *f* X) whn" | 
| 15 | by (simp add: hypnat_omega_def starfun_def star_of_def Ifun_star_n) | |
| 16 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 17 | lemma starset_n_Un: "*sn* (%n. (A n) Un (B n)) = *sn* A Un *sn* B" | 
| 17443 | 18 | apply (simp add: starset_n_def star_n_eq_starfun_whn Un_def) | 
| 19 | apply (rule_tac x=whn in spec, transfer, simp) | |
| 14415 | 20 | done | 
| 21 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 22 | lemma InternalSets_Un: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 23 | "[| X \<in> InternalSets; Y \<in> InternalSets |] | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 24 | ==> (X Un Y) \<in> InternalSets" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 25 | by (auto simp add: InternalSets_def starset_n_Un [symmetric]) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 26 | |
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 27 | lemma starset_n_Int: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 28 | "*sn* (%n. (A n) Int (B n)) = *sn* A Int *sn* B" | 
| 17443 | 29 | apply (simp add: starset_n_def star_n_eq_starfun_whn Int_def) | 
| 30 | apply (rule_tac x=whn in spec, transfer, simp) | |
| 14415 | 31 | done | 
| 32 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 33 | lemma InternalSets_Int: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 34 | "[| X \<in> InternalSets; Y \<in> InternalSets |] | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 35 | ==> (X Int Y) \<in> InternalSets" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 36 | by (auto simp add: InternalSets_def starset_n_Int [symmetric]) | 
| 14415 | 37 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 38 | lemma starset_n_Compl: "*sn* ((%n. - A n)) = -( *sn* A)" | 
| 17443 | 39 | apply (simp add: starset_n_def star_n_eq_starfun_whn Compl_def) | 
| 40 | apply (rule_tac x=whn in spec, transfer, simp) | |
| 14415 | 41 | done | 
| 42 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 43 | lemma InternalSets_Compl: "X \<in> InternalSets ==> -X \<in> InternalSets" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 44 | by (auto simp add: InternalSets_def starset_n_Compl [symmetric]) | 
| 14415 | 45 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 46 | lemma starset_n_diff: "*sn* (%n. (A n) - (B n)) = *sn* A - *sn* B" | 
| 17443 | 47 | apply (simp add: starset_n_def star_n_eq_starfun_whn set_diff_def) | 
| 48 | apply (rule_tac x=whn in spec, transfer, simp) | |
| 14415 | 49 | done | 
| 50 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 51 | lemma InternalSets_diff: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 52 | "[| X \<in> InternalSets; Y \<in> InternalSets |] | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 53 | ==> (X - Y) \<in> InternalSets" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 54 | by (auto simp add: InternalSets_def starset_n_diff [symmetric]) | 
| 14415 | 55 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 56 | lemma NatStar_SHNat_subset: "Nats \<le> *s* (UNIV:: nat set)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 57 | by simp | 
| 14415 | 58 | |
| 59 | lemma NatStar_hypreal_of_real_Int: | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 60 | "*s* X Int Nats = hypnat_of_nat ` X" | 
| 20732 | 61 | by (auto simp add: SHNat_eq) | 
| 14415 | 62 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 63 | lemma starset_starset_n_eq: "*s* X = *sn* (%n. X)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 64 | by (simp add: starset_n_starset) | 
| 14415 | 65 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 66 | lemma InternalSets_starset_n [simp]: "( *s* X) \<in> InternalSets" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 67 | by (auto simp add: InternalSets_def starset_starset_n_eq) | 
| 14415 | 68 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 69 | lemma InternalSets_UNIV_diff: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 70 | "X \<in> InternalSets ==> UNIV - X \<in> InternalSets" | 
| 17290 | 71 | apply (subgoal_tac "UNIV - X = - X") | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 72 | by (auto intro: InternalSets_Compl) | 
| 14415 | 73 | |
| 74 | ||
| 75 | subsection{*Nonstandard Extensions of Functions*}
 | |
| 76 | ||
| 77 | text{* Example of transfer of a property from reals to hyperreals
 | |
| 78 | --- used for limit comparison of sequences*} | |
| 79 | ||
| 80 | lemma starfun_le_mono: | |
| 81 | "\<forall>n. N \<le> n --> f n \<le> g n | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 82 | ==> \<forall>n. hypnat_of_nat N \<le> n --> ( *f* f) n \<le> ( *f* g) n" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 83 | by transfer | 
| 14415 | 84 | |
| 85 | (*****----- and another -----*****) | |
| 86 | lemma starfun_less_mono: | |
| 87 | "\<forall>n. N \<le> n --> f n < g n | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 88 | ==> \<forall>n. hypnat_of_nat N \<le> n --> ( *f* f) n < ( *f* g) n" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 89 | by transfer | 
| 14415 | 90 | |
| 91 | text{*Nonstandard extension when we increment the argument by one*}
 | |
| 92 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 93 | lemma starfun_shift_one: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 94 | "!!N. ( *f* (%n. f (Suc n))) N = ( *f* f) (N + (1::hypnat))" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 95 | by (transfer, simp) | 
| 14415 | 96 | |
| 97 | text{*Nonstandard extension with absolute value*}
 | |
| 98 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 99 | lemma starfun_abs: "!!N. ( *f* (%n. abs (f n))) N = abs(( *f* f) N)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 100 | by (transfer, rule refl) | 
| 14415 | 101 | |
| 102 | text{*The hyperpow function as a nonstandard extension of realpow*}
 | |
| 103 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 104 | lemma starfun_pow: "!!N. ( *f* (%n. r ^ n)) N = (hypreal_of_real r) pow N" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 105 | by (transfer, rule refl) | 
| 14415 | 106 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 107 | lemma starfun_pow2: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 108 | "!!N. ( *f* (%n. (X n) ^ m)) N = ( *f* X) N pow hypnat_of_nat m" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 109 | by (transfer, rule refl) | 
| 14415 | 110 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 111 | lemma starfun_pow3: "!!R. ( *f* (%r. r ^ n)) R = (R) pow hypnat_of_nat n" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 112 | by (transfer, rule refl) | 
| 14415 | 113 | |
| 114 | text{*The @{term hypreal_of_hypnat} function as a nonstandard extension of
 | |
| 115 |   @{term real_of_nat} *}
 | |
| 116 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 117 | lemma starfunNat_real_of_nat: "( *f* real) = hypreal_of_hypnat" | 
| 21864 
2ecfd8985982
hypreal_of_hypnat abbreviates more general of_hypnat
 huffman parents: 
21847diff
changeset | 118 | by transfer (simp add: expand_fun_eq real_of_nat_def) | 
| 14415 | 119 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 120 | lemma starfun_inverse_real_of_nat_eq: | 
| 14415 | 121 | "N \<in> HNatInfinite | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 122 | ==> ( *f* (%x::nat. inverse(real x))) N = inverse(hypreal_of_hypnat N)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 123 | apply (rule_tac f1 = inverse in starfun_o2 [THEN subst]) | 
| 14415 | 124 | apply (subgoal_tac "hypreal_of_hypnat N ~= 0") | 
| 20740 
5a103b43da5a
reorganized HNatInfinite proofs; simplified and renamed some lemmas
 huffman parents: 
20732diff
changeset | 125 | apply (simp_all add: zero_less_HNatInfinite starfunNat_real_of_nat starfun_inverse_inverse) | 
| 14415 | 126 | done | 
| 127 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 128 | text{*Internal functions - some redundancy with *f* now*}
 | 
| 14415 | 129 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 130 | lemma starfun_n: "( *fn* f) (star_n X) = star_n (%n. f n (X n))" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 131 | by (simp add: starfun_n_def Ifun_star_n) | 
| 14415 | 132 | |
| 133 | text{*Multiplication: @{text "( *fn) x ( *gn) = *(fn x gn)"}*}
 | |
| 10751 | 134 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 135 | lemma starfun_n_mult: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 136 | "( *fn* f) z * ( *fn* g) z = ( *fn* (% i x. f i x * g i x)) z" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 137 | apply (cases z) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 138 | apply (simp add: starfun_n star_n_mult) | 
| 14415 | 139 | done | 
| 140 | ||
| 141 | text{*Addition: @{text "( *fn) + ( *gn) = *(fn + gn)"}*}
 | |
| 142 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 143 | lemma starfun_n_add: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 144 | "( *fn* f) z + ( *fn* g) z = ( *fn* (%i x. f i x + g i x)) z" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 145 | apply (cases z) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 146 | apply (simp add: starfun_n star_n_add) | 
| 14415 | 147 | done | 
| 148 | ||
| 149 | text{*Subtraction: @{text "( *fn) - ( *gn) = *(fn + - gn)"}*}
 | |
| 150 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 151 | lemma starfun_n_add_minus: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 152 | "( *fn* f) z + -( *fn* g) z = ( *fn* (%i x. f i x + -g i x)) z" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 153 | apply (cases z) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 154 | apply (simp add: starfun_n star_n_minus star_n_add) | 
| 14415 | 155 | done | 
| 156 | ||
| 157 | ||
| 158 | text{*Composition: @{text "( *fn) o ( *gn) = *(fn o gn)"}*}
 | |
| 159 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 160 | lemma starfun_n_const_fun [simp]: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 161 | "( *fn* (%i x. k)) z = star_of k" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 162 | apply (cases z) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 163 | apply (simp add: starfun_n star_of_def) | 
| 14415 | 164 | done | 
| 165 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 166 | lemma starfun_n_minus: "- ( *fn* f) x = ( *fn* (%i x. - (f i) x)) x" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 167 | apply (cases x) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 168 | apply (simp add: starfun_n star_n_minus) | 
| 14415 | 169 | done | 
| 170 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 171 | lemma starfun_n_eq [simp]: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 172 | "( *fn* f) (star_of n) = star_n (%i. f i n)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 173 | by (simp add: starfun_n star_of_def) | 
| 14415 | 174 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 175 | lemma starfun_eq_iff: "(( *f* f) = ( *f* g)) = (f = g)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 176 | by (transfer, rule refl) | 
| 14415 | 177 | |
| 178 | lemma starfunNat_inverse_real_of_nat_Infinitesimal [simp]: | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 179 | "N \<in> HNatInfinite ==> ( *f* (%x. inverse (real x))) N \<in> Infinitesimal" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 180 | apply (rule_tac f1 = inverse in starfun_o2 [THEN subst]) | 
| 14415 | 181 | apply (subgoal_tac "hypreal_of_hypnat N ~= 0") | 
| 20740 
5a103b43da5a
reorganized HNatInfinite proofs; simplified and renamed some lemmas
 huffman parents: 
20732diff
changeset | 182 | apply (simp_all add: zero_less_HNatInfinite starfunNat_real_of_nat) | 
| 14415 | 183 | done | 
| 184 | ||
| 14641 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 185 | |
| 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 186 | subsection{*Nonstandard Characterization of Induction*}
 | 
| 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 187 | |
| 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 188 | lemma hypnat_induct_obj: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 189 | "!!n. (( *p* P) (0::hypnat) & | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 190 | (\<forall>n. ( *p* P)(n) --> ( *p* P)(n + 1))) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 191 | --> ( *p* P)(n)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 192 | by (transfer, induct_tac n, auto) | 
| 14641 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 193 | |
| 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 194 | lemma hypnat_induct: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 195 | "!!n. [| ( *p* P) (0::hypnat); | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 196 | !!n. ( *p* P)(n) ==> ( *p* P)(n + 1)|] | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 197 | ==> ( *p* P)(n)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 198 | by (transfer, induct_tac n, auto) | 
| 14641 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 199 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 200 | lemma starP2_eq_iff: "( *p2* (op =)) = (op =)" | 
| 21847 
59a68ed9f2f2
redefine hSuc as *f* Suc, and move to HyperNat.thy
 huffman parents: 
21404diff
changeset | 201 | by transfer (rule refl) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 202 | |
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 203 | lemma starP2_eq_iff2: "( *p2* (%x y. x = y)) X Y = (X = Y)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 204 | by (simp add: starP2_eq_iff) | 
| 14641 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 205 | |
| 21847 
59a68ed9f2f2
redefine hSuc as *f* Suc, and move to HyperNat.thy
 huffman parents: 
21404diff
changeset | 206 | lemma nonempty_nat_set_Least_mem: | 
| 
59a68ed9f2f2
redefine hSuc as *f* Suc, and move to HyperNat.thy
 huffman parents: 
21404diff
changeset | 207 | "c \<in> (S :: nat set) ==> (LEAST n. n \<in> S) \<in> S" | 
| 14641 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 208 | by (erule LeastI) | 
| 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 209 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 210 | lemma nonempty_set_star_has_least: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 211 |     "!!S::nat set star. Iset S \<noteq> {} ==> \<exists>n \<in> Iset S. \<forall>m \<in> Iset S. n \<le> m"
 | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 212 | apply (transfer empty_def) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 213 | apply (rule_tac x="LEAST n. n \<in> S" in bexI) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 214 | apply (simp add: Least_le) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 215 | apply (rule LeastI_ex, auto) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 216 | done | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 217 | |
| 14641 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 218 | lemma nonempty_InternalNatSet_has_least: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 219 |     "[| (S::hypnat set) \<in> InternalSets; S \<noteq> {} |] ==> \<exists>n \<in> S. \<forall>m \<in> S. n \<le> m"
 | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 220 | apply (clarsimp simp add: InternalSets_def starset_n_def) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 221 | apply (erule nonempty_set_star_has_least) | 
| 14641 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 222 | done | 
| 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 223 | |
| 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 224 | text{* Goldblatt page 129 Thm 11.3.2*}
 | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 225 | lemma internal_induct_lemma: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 226 | "!!X::nat set star. [| (0::hypnat) \<in> Iset X; \<forall>n. n \<in> Iset X --> n + 1 \<in> Iset X |] | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 227 | ==> Iset X = (UNIV:: hypnat set)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 228 | apply (transfer UNIV_def) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 229 | apply (rule equalityI [OF subset_UNIV subsetI]) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 230 | apply (induct_tac x, auto) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 231 | done | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 232 | |
| 14641 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 233 | lemma internal_induct: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 234 | "[| X \<in> InternalSets; (0::hypnat) \<in> X; \<forall>n. n \<in> X --> n + 1 \<in> X |] | 
| 14641 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 235 | ==> X = (UNIV:: hypnat set)" | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 236 | apply (clarsimp simp add: InternalSets_def starset_n_def) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17299diff
changeset | 237 | apply (erule (1) internal_induct_lemma) | 
| 14641 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 238 | done | 
| 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 239 | |
| 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 paulson parents: 
14468diff
changeset | 240 | |
| 10751 | 241 | end |