| author | nipkow | 
| Fri, 15 Apr 2005 18:43:35 +0200 | |
| changeset 15740 | d63e7a65b2d0 | 
| parent 14153 | 76a6ba67bd15 | 
| child 16417 | 9bc16273c2d4 | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/Nat.thy | 
| 0 | 2 | ID: $Id$ | 
| 1478 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 435 | 4 | Copyright 1994 University of Cambridge | 
| 0 | 5 | |
| 6 | *) | |
| 7 | ||
| 13356 | 8 | header{*The Natural numbers As a Least Fixed Point*}
 | 
| 9 | ||
| 13357 | 10 | theory Nat = OrdQuant + Bool: | 
| 0 | 11 | |
| 12789 | 12 | constdefs | 
| 13 | nat :: i | |
| 14 |     "nat == lfp(Inf, %X. {0} Un {succ(i). i:X})"
 | |
| 0 | 15 | |
| 13269 | 16 | quasinat :: "i => o" | 
| 17 | "quasinat(n) == n=0 | (\<exists>m. n = succ(m))" | |
| 18 | ||
| 13173 | 19 | (*Has an unconditional succ case, which is used in "recursor" below.*) | 
| 12789 | 20 | nat_case :: "[i, i=>i, i]=>i" | 
| 21 | "nat_case(a,b,k) == THE y. k=0 & y=a | (EX x. k=succ(x) & y=b(x))" | |
| 0 | 22 | |
| 12789 | 23 | nat_rec :: "[i, i, [i,i]=>i]=>i" | 
| 24 | "nat_rec(k,a,b) == | |
| 1478 | 25 | wfrec(Memrel(nat), k, %n f. nat_case(a, %m. b(m, f`m), n))" | 
| 0 | 26 | |
| 12789 | 27 | (*Internalized relations on the naturals*) | 
| 28 | ||
| 29 | Le :: i | |
| 30 |     "Le == {<x,y>:nat*nat. x le y}"
 | |
| 31 | ||
| 32 | Lt :: i | |
| 33 |     "Lt == {<x, y>:nat*nat. x < y}"
 | |
| 34 | ||
| 35 | Ge :: i | |
| 36 |     "Ge == {<x,y>:nat*nat. y le x}"
 | |
| 37 | ||
| 38 | Gt :: i | |
| 39 |     "Gt == {<x,y>:nat*nat. y < x}"
 | |
| 40 | ||
| 13171 | 41 | greater_than :: "i=>i" | 
| 12789 | 42 |     "greater_than(n) == {i:nat. n < i}"
 | 
| 43 | ||
| 14046 | 44 | text{*No need for a less-than operator: a natural number is its list of
 | 
| 45 | predecessors!*} | |
| 46 | ||
| 47 | ||
| 13171 | 48 | lemma nat_bnd_mono: "bnd_mono(Inf, %X. {0} Un {succ(i). i:X})"
 | 
| 49 | apply (rule bnd_monoI) | |
| 13269 | 50 | apply (cut_tac infinity, blast, blast) | 
| 13171 | 51 | done | 
| 52 | ||
| 53 | (* nat = {0} Un {succ(x). x:nat} *)
 | |
| 54 | lemmas nat_unfold = nat_bnd_mono [THEN nat_def [THEN def_lfp_unfold], standard] | |
| 55 | ||
| 56 | (** Type checking of 0 and successor **) | |
| 57 | ||
| 58 | lemma nat_0I [iff,TC]: "0 : nat" | |
| 59 | apply (subst nat_unfold) | |
| 60 | apply (rule singletonI [THEN UnI1]) | |
| 61 | done | |
| 62 | ||
| 63 | lemma nat_succI [intro!,TC]: "n : nat ==> succ(n) : nat" | |
| 64 | apply (subst nat_unfold) | |
| 65 | apply (erule RepFunI [THEN UnI2]) | |
| 66 | done | |
| 67 | ||
| 68 | lemma nat_1I [iff,TC]: "1 : nat" | |
| 69 | by (rule nat_0I [THEN nat_succI]) | |
| 70 | ||
| 71 | lemma nat_2I [iff,TC]: "2 : nat" | |
| 72 | by (rule nat_1I [THEN nat_succI]) | |
| 73 | ||
| 74 | lemma bool_subset_nat: "bool <= nat" | |
| 75 | by (blast elim!: boolE) | |
| 76 | ||
| 77 | lemmas bool_into_nat = bool_subset_nat [THEN subsetD, standard] | |
| 78 | ||
| 79 | ||
| 13356 | 80 | subsection{*Injectivity Properties and Induction*}
 | 
| 13171 | 81 | |
| 82 | (*Mathematical induction*) | |
| 13524 | 83 | lemma nat_induct [case_names 0 succ, induct set: nat]: | 
| 13171 | 84 | "[| n: nat; P(0); !!x. [| x: nat; P(x) |] ==> P(succ(x)) |] ==> P(n)" | 
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 85 | by (erule def_induct [OF nat_def nat_bnd_mono], blast) | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 86 | |
| 13171 | 87 | lemma natE: | 
| 88 | "[| n: nat; n=0 ==> P; !!x. [| x: nat; n=succ(x) |] ==> P |] ==> P" | |
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 89 | by (erule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE], auto) | 
| 13171 | 90 | |
| 91 | lemma nat_into_Ord [simp]: "n: nat ==> Ord(n)" | |
| 92 | by (erule nat_induct, auto) | |
| 93 | ||
| 94 | (* i: nat ==> 0 le i; same thing as 0<succ(i) *) | |
| 95 | lemmas nat_0_le = nat_into_Ord [THEN Ord_0_le, standard] | |
| 96 | ||
| 97 | (* i: nat ==> i le i; same thing as i<succ(i) *) | |
| 98 | lemmas nat_le_refl = nat_into_Ord [THEN le_refl, standard] | |
| 99 | ||
| 100 | lemma Ord_nat [iff]: "Ord(nat)" | |
| 101 | apply (rule OrdI) | |
| 102 | apply (erule_tac [2] nat_into_Ord [THEN Ord_is_Transset]) | |
| 103 | apply (unfold Transset_def) | |
| 104 | apply (rule ballI) | |
| 105 | apply (erule nat_induct, auto) | |
| 106 | done | |
| 107 | ||
| 108 | lemma Limit_nat [iff]: "Limit(nat)" | |
| 109 | apply (unfold Limit_def) | |
| 110 | apply (safe intro!: ltI Ord_nat) | |
| 111 | apply (erule ltD) | |
| 112 | done | |
| 113 | ||
| 13628 
87482b5e3f2e
Various simplifications of the Constructible theories
 paulson parents: 
13524diff
changeset | 114 | lemma naturals_not_limit: "a \<in> nat ==> ~ Limit(a)" | 
| 
87482b5e3f2e
Various simplifications of the Constructible theories
 paulson parents: 
13524diff
changeset | 115 | by (induct a rule: nat_induct, auto) | 
| 
87482b5e3f2e
Various simplifications of the Constructible theories
 paulson parents: 
13524diff
changeset | 116 | |
| 13823 | 117 | lemma succ_natD: "succ(i): nat ==> i: nat" | 
| 13171 | 118 | by (rule Ord_trans [OF succI1], auto) | 
| 119 | ||
| 120 | lemma nat_succ_iff [iff]: "succ(n): nat <-> n: nat" | |
| 13823 | 121 | by (blast dest!: succ_natD) | 
| 13171 | 122 | |
| 123 | lemma nat_le_Limit: "Limit(i) ==> nat le i" | |
| 124 | apply (rule subset_imp_le) | |
| 125 | apply (simp_all add: Limit_is_Ord) | |
| 126 | apply (rule subsetI) | |
| 127 | apply (erule nat_induct) | |
| 128 | apply (erule Limit_has_0 [THEN ltD]) | |
| 129 | apply (blast intro: Limit_has_succ [THEN ltD] ltI Limit_is_Ord) | |
| 130 | done | |
| 131 | ||
| 132 | (* [| succ(i): k; k: nat |] ==> i: k *) | |
| 133 | lemmas succ_in_naturalD = Ord_trans [OF succI1 _ nat_into_Ord] | |
| 134 | ||
| 135 | lemma lt_nat_in_nat: "[| m<n; n: nat |] ==> m: nat" | |
| 136 | apply (erule ltE) | |
| 13269 | 137 | apply (erule Ord_trans, assumption, simp) | 
| 13171 | 138 | done | 
| 139 | ||
| 140 | lemma le_in_nat: "[| m le n; n:nat |] ==> m:nat" | |
| 141 | by (blast dest!: lt_nat_in_nat) | |
| 142 | ||
| 143 | ||
| 13356 | 144 | subsection{*Variations on Mathematical Induction*}
 | 
| 13171 | 145 | |
| 146 | (*complete induction*) | |
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 147 | |
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 148 | lemmas complete_induct = Ord_induct [OF _ Ord_nat, case_names less, consumes 1] | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 149 | |
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 150 | lemmas complete_induct_rule = | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 151 | complete_induct [rule_format, case_names less, consumes 1] | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 152 | |
| 13171 | 153 | |
| 154 | lemma nat_induct_from_lemma [rule_format]: | |
| 155 | "[| n: nat; m: nat; | |
| 156 | !!x. [| x: nat; m le x; P(x) |] ==> P(succ(x)) |] | |
| 157 | ==> m le n --> P(m) --> P(n)" | |
| 158 | apply (erule nat_induct) | |
| 159 | apply (simp_all add: distrib_simps le0_iff le_succ_iff) | |
| 160 | done | |
| 161 | ||
| 162 | (*Induction starting from m rather than 0*) | |
| 163 | lemma nat_induct_from: | |
| 164 | "[| m le n; m: nat; n: nat; | |
| 165 | P(m); | |
| 166 | !!x. [| x: nat; m le x; P(x) |] ==> P(succ(x)) |] | |
| 167 | ==> P(n)" | |
| 168 | apply (blast intro: nat_induct_from_lemma) | |
| 169 | done | |
| 170 | ||
| 171 | (*Induction suitable for subtraction and less-than*) | |
| 13524 | 172 | lemma diff_induct [case_names 0 0_succ succ_succ, consumes 2]: | 
| 13171 | 173 | "[| m: nat; n: nat; | 
| 174 | !!x. x: nat ==> P(x,0); | |
| 175 | !!y. y: nat ==> P(0,succ(y)); | |
| 176 | !!x y. [| x: nat; y: nat; P(x,y) |] ==> P(succ(x),succ(y)) |] | |
| 177 | ==> P(m,n)" | |
| 13784 | 178 | apply (erule_tac x = m in rev_bspec) | 
| 13171 | 179 | apply (erule nat_induct, simp) | 
| 180 | apply (rule ballI) | |
| 181 | apply (rename_tac i j) | |
| 182 | apply (erule_tac n=j in nat_induct, auto) | |
| 183 | done | |
| 184 | ||
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 185 | |
| 13171 | 186 | (** Induction principle analogous to trancl_induct **) | 
| 187 | ||
| 188 | lemma succ_lt_induct_lemma [rule_format]: | |
| 189 | "m: nat ==> P(m,succ(m)) --> (ALL x: nat. P(m,x) --> P(m,succ(x))) --> | |
| 190 | (ALL n:nat. m<n --> P(m,n))" | |
| 191 | apply (erule nat_induct) | |
| 192 | apply (intro impI, rule nat_induct [THEN ballI]) | |
| 193 | prefer 4 apply (intro impI, rule nat_induct [THEN ballI]) | |
| 194 | apply (auto simp add: le_iff) | |
| 195 | done | |
| 196 | ||
| 197 | lemma succ_lt_induct: | |
| 198 | "[| m<n; n: nat; | |
| 199 | P(m,succ(m)); | |
| 200 | !!x. [| x: nat; P(m,x) |] ==> P(m,succ(x)) |] | |
| 201 | ==> P(m,n)" | |
| 202 | by (blast intro: succ_lt_induct_lemma lt_nat_in_nat) | |
| 203 | ||
| 13269 | 204 | subsection{*quasinat: to allow a case-split rule for @{term nat_case}*}
 | 
| 205 | ||
| 206 | text{*True if the argument is zero or any successor*}
 | |
| 207 | lemma [iff]: "quasinat(0)" | |
| 208 | by (simp add: quasinat_def) | |
| 209 | ||
| 210 | lemma [iff]: "quasinat(succ(x))" | |
| 211 | by (simp add: quasinat_def) | |
| 212 | ||
| 213 | lemma nat_imp_quasinat: "n \<in> nat ==> quasinat(n)" | |
| 214 | by (erule natE, simp_all) | |
| 215 | ||
| 216 | lemma non_nat_case: "~ quasinat(x) ==> nat_case(a,b,x) = 0" | |
| 217 | by (simp add: quasinat_def nat_case_def) | |
| 218 | ||
| 219 | lemma nat_cases_disj: "k=0 | (\<exists>y. k = succ(y)) | ~ quasinat(k)" | |
| 14153 | 220 | apply (case_tac "k=0", simp) | 
| 221 | apply (case_tac "\<exists>m. k = succ(m)") | |
| 222 | apply (simp_all add: quasinat_def) | |
| 13269 | 223 | done | 
| 224 | ||
| 225 | lemma nat_cases: | |
| 226 | "[|k=0 ==> P; !!y. k = succ(y) ==> P; ~ quasinat(k) ==> P|] ==> P" | |
| 13356 | 227 | by (insert nat_cases_disj [of k], blast) | 
| 13269 | 228 | |
| 13171 | 229 | (** nat_case **) | 
| 230 | ||
| 231 | lemma nat_case_0 [simp]: "nat_case(a,b,0) = a" | |
| 13174 | 232 | by (simp add: nat_case_def) | 
| 13173 | 233 | |
| 234 | lemma nat_case_succ [simp]: "nat_case(a,b,succ(n)) = b(n)" | |
| 13174 | 235 | by (simp add: nat_case_def) | 
| 13171 | 236 | |
| 13173 | 237 | lemma nat_case_type [TC]: | 
| 238 | "[| n: nat; a: C(0); !!m. m: nat ==> b(m): C(succ(m)) |] | |
| 239 | ==> nat_case(a,b,n) : C(n)"; | |
| 240 | by (erule nat_induct, auto) | |
| 241 | ||
| 13269 | 242 | lemma split_nat_case: | 
| 243 | "P(nat_case(a,b,k)) <-> | |
| 244 | ((k=0 --> P(a)) & (\<forall>x. k=succ(x) --> P(b(x))) & (~ quasinat(k) \<longrightarrow> P(0)))" | |
| 245 | apply (rule nat_cases [of k]) | |
| 246 | apply (auto simp add: non_nat_case) | |
| 13171 | 247 | done | 
| 248 | ||
| 13173 | 249 | |
| 13356 | 250 | subsection{*Recursion on the Natural Numbers*}
 | 
| 13171 | 251 | |
| 13356 | 252 | (** nat_rec is used to define eclose and transrec, then becomes obsolete. | 
| 253 | The operator rec, from arith.thy, has fewer typing conditions **) | |
| 13171 | 254 | |
| 255 | lemma nat_rec_0: "nat_rec(0,a,b) = a" | |
| 256 | apply (rule nat_rec_def [THEN def_wfrec, THEN trans]) | |
| 257 | apply (rule wf_Memrel) | |
| 258 | apply (rule nat_case_0) | |
| 259 | done | |
| 260 | ||
| 261 | lemma nat_rec_succ: "m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))" | |
| 262 | apply (rule nat_rec_def [THEN def_wfrec, THEN trans]) | |
| 263 | apply (rule wf_Memrel) | |
| 264 | apply (simp add: vimage_singleton_iff) | |
| 265 | done | |
| 266 | ||
| 267 | (** The union of two natural numbers is a natural number -- their maximum **) | |
| 268 | ||
| 13173 | 269 | lemma Un_nat_type [TC]: "[| i: nat; j: nat |] ==> i Un j: nat" | 
| 13171 | 270 | apply (rule Un_least_lt [THEN ltD]) | 
| 271 | apply (simp_all add: lt_def) | |
| 272 | done | |
| 273 | ||
| 13173 | 274 | lemma Int_nat_type [TC]: "[| i: nat; j: nat |] ==> i Int j: nat" | 
| 13171 | 275 | apply (rule Int_greatest_lt [THEN ltD]) | 
| 276 | apply (simp_all add: lt_def) | |
| 277 | done | |
| 278 | ||
| 279 | (*needed to simplify unions over nat*) | |
| 280 | lemma nat_nonempty [simp]: "nat ~= 0" | |
| 281 | by blast | |
| 282 | ||
| 14046 | 283 | text{*A natural number is the set of its predecessors*}
 | 
| 284 | lemma nat_eq_Collect_lt: "i \<in> nat ==> {j\<in>nat. j<i} = i"
 | |
| 285 | apply (rule equalityI) | |
| 286 | apply (blast dest: ltD) | |
| 287 | apply (auto simp add: Ord_mem_iff_lt) | |
| 288 | apply (blast intro: lt_trans) | |
| 289 | done | |
| 290 | ||
| 14054 | 291 | lemma Le_iff [iff]: "<x,y> : Le <-> x le y & x : nat & y : nat" | 
| 292 | by (force simp add: Le_def) | |
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 293 | |
| 13171 | 294 | ML | 
| 295 | {*
 | |
| 296 | val Le_def = thm "Le_def"; | |
| 297 | val Lt_def = thm "Lt_def"; | |
| 298 | val Ge_def = thm "Ge_def"; | |
| 299 | val Gt_def = thm "Gt_def"; | |
| 300 | val greater_than_def = thm "greater_than_def"; | |
| 301 | ||
| 302 | val nat_bnd_mono = thm "nat_bnd_mono"; | |
| 303 | val nat_unfold = thm "nat_unfold"; | |
| 304 | val nat_0I = thm "nat_0I"; | |
| 305 | val nat_succI = thm "nat_succI"; | |
| 306 | val nat_1I = thm "nat_1I"; | |
| 307 | val nat_2I = thm "nat_2I"; | |
| 308 | val bool_subset_nat = thm "bool_subset_nat"; | |
| 309 | val bool_into_nat = thm "bool_into_nat"; | |
| 310 | val nat_induct = thm "nat_induct"; | |
| 311 | val natE = thm "natE"; | |
| 312 | val nat_into_Ord = thm "nat_into_Ord"; | |
| 313 | val nat_0_le = thm "nat_0_le"; | |
| 314 | val nat_le_refl = thm "nat_le_refl"; | |
| 315 | val Ord_nat = thm "Ord_nat"; | |
| 316 | val Limit_nat = thm "Limit_nat"; | |
| 317 | val succ_natD = thm "succ_natD"; | |
| 318 | val nat_succ_iff = thm "nat_succ_iff"; | |
| 319 | val nat_le_Limit = thm "nat_le_Limit"; | |
| 320 | val succ_in_naturalD = thm "succ_in_naturalD"; | |
| 321 | val lt_nat_in_nat = thm "lt_nat_in_nat"; | |
| 322 | val le_in_nat = thm "le_in_nat"; | |
| 323 | val complete_induct = thm "complete_induct"; | |
| 324 | val nat_induct_from = thm "nat_induct_from"; | |
| 325 | val diff_induct = thm "diff_induct"; | |
| 326 | val succ_lt_induct = thm "succ_lt_induct"; | |
| 327 | val nat_case_0 = thm "nat_case_0"; | |
| 328 | val nat_case_succ = thm "nat_case_succ"; | |
| 329 | val nat_case_type = thm "nat_case_type"; | |
| 330 | val nat_rec_0 = thm "nat_rec_0"; | |
| 331 | val nat_rec_succ = thm "nat_rec_succ"; | |
| 332 | val Un_nat_type = thm "Un_nat_type"; | |
| 333 | val Int_nat_type = thm "Int_nat_type"; | |
| 334 | val nat_nonempty = thm "nat_nonempty"; | |
| 335 | *} | |
| 336 | ||
| 0 | 337 | end |