author | paulson |
Sat, 22 Dec 2001 19:46:16 +0100 | |
changeset 12594 | 5b9b0adca8aa |
parent 12485 | 3df60299a6d0 |
child 12825 | f1f7964ed05c |
permissions | -rw-r--r-- |
9907 | 1 |
(* Title: ZF/Ordinal.ML |
435 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
435 | 4 |
Copyright 1993 University of Cambridge |
5 |
||
5321 | 6 |
Ordinals in Zermelo-Fraenkel Set Theory |
435 | 7 |
*) |
8 |
||
9 |
(*** Rules for Transset ***) |
|
10 |
||
11076 | 11 |
(** Three neat characterisations of Transset **) |
435 | 12 |
|
5067 | 13 |
Goalw [Transset_def] "Transset(A) <-> A<=Pow(A)"; |
2925 | 14 |
by (Blast_tac 1); |
760 | 15 |
qed "Transset_iff_Pow"; |
435 | 16 |
|
5067 | 17 |
Goalw [Transset_def] "Transset(A) <-> Union(succ(A)) = A"; |
4091 | 18 |
by (blast_tac (claset() addSEs [equalityE]) 1); |
760 | 19 |
qed "Transset_iff_Union_succ"; |
435 | 20 |
|
11076 | 21 |
Goalw [Transset_def] "Transset(A) <-> Union(A) <= A"; |
22 |
by (Blast_tac 1); |
|
23 |
qed "Transset_iff_Union_subset"; |
|
24 |
||
435 | 25 |
(** Consequences of downwards closure **) |
26 |
||
5067 | 27 |
Goalw [Transset_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
28 |
"[| Transset(C); {a,b}: C |] ==> a:C & b: C"; |
2925 | 29 |
by (Blast_tac 1); |
760 | 30 |
qed "Transset_doubleton_D"; |
435 | 31 |
|
9907 | 32 |
val [prem1,prem2] = goalw (the_context ()) [Pair_def] |
435 | 33 |
"[| Transset(C); <a,b>: C |] ==> a:C & b: C"; |
1461 | 34 |
by (cut_facts_tac [prem2] 1); |
4091 | 35 |
by (blast_tac (claset() addSDs [prem1 RS Transset_doubleton_D]) 1); |
760 | 36 |
qed "Transset_Pair_D"; |
435 | 37 |
|
9907 | 38 |
val prem1::prems = goal (the_context ()) |
435 | 39 |
"[| Transset(C); A*B <= C; b: B |] ==> A <= C"; |
40 |
by (cut_facts_tac prems 1); |
|
4091 | 41 |
by (blast_tac (claset() addSDs [prem1 RS Transset_Pair_D]) 1); |
760 | 42 |
qed "Transset_includes_domain"; |
435 | 43 |
|
9907 | 44 |
val prem1::prems = goal (the_context ()) |
435 | 45 |
"[| Transset(C); A*B <= C; a: A |] ==> B <= C"; |
46 |
by (cut_facts_tac prems 1); |
|
4091 | 47 |
by (blast_tac (claset() addSDs [prem1 RS Transset_Pair_D]) 1); |
760 | 48 |
qed "Transset_includes_range"; |
435 | 49 |
|
50 |
(** Closure properties **) |
|
51 |
||
5067 | 52 |
Goalw [Transset_def] "Transset(0)"; |
2925 | 53 |
by (Blast_tac 1); |
760 | 54 |
qed "Transset_0"; |
435 | 55 |
|
5067 | 56 |
Goalw [Transset_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
57 |
"[| Transset(i); Transset(j) |] ==> Transset(i Un j)"; |
2925 | 58 |
by (Blast_tac 1); |
760 | 59 |
qed "Transset_Un"; |
435 | 60 |
|
5067 | 61 |
Goalw [Transset_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
62 |
"[| Transset(i); Transset(j) |] ==> Transset(i Int j)"; |
2925 | 63 |
by (Blast_tac 1); |
760 | 64 |
qed "Transset_Int"; |
435 | 65 |
|
5137 | 66 |
Goalw [Transset_def] "Transset(i) ==> Transset(succ(i))"; |
2925 | 67 |
by (Blast_tac 1); |
760 | 68 |
qed "Transset_succ"; |
435 | 69 |
|
5137 | 70 |
Goalw [Transset_def] "Transset(i) ==> Transset(Pow(i))"; |
2925 | 71 |
by (Blast_tac 1); |
760 | 72 |
qed "Transset_Pow"; |
435 | 73 |
|
5137 | 74 |
Goalw [Transset_def] "Transset(A) ==> Transset(Union(A))"; |
2925 | 75 |
by (Blast_tac 1); |
760 | 76 |
qed "Transset_Union"; |
435 | 77 |
|
5321 | 78 |
val [Transprem] = Goalw [Transset_def] |
435 | 79 |
"[| !!i. i:A ==> Transset(i) |] ==> Transset(Union(A))"; |
4091 | 80 |
by (blast_tac (claset() addDs [Transprem RS bspec RS subsetD]) 1); |
760 | 81 |
qed "Transset_Union_family"; |
435 | 82 |
|
5321 | 83 |
val [prem,Transprem] = Goalw [Transset_def] |
435 | 84 |
"[| j:A; !!i. i:A ==> Transset(i) |] ==> Transset(Inter(A))"; |
85 |
by (cut_facts_tac [prem] 1); |
|
4091 | 86 |
by (blast_tac (claset() addDs [Transprem RS bspec RS subsetD]) 1); |
760 | 87 |
qed "Transset_Inter_family"; |
435 | 88 |
|
89 |
(*** Natural Deduction rules for Ord ***) |
|
90 |
||
5321 | 91 |
val prems = Goalw [Ord_def] |
2717
b29c45ef3d86
best_tac avoids looping with change to RepFun_eqI in claset
paulson
parents:
2493
diff
changeset
|
92 |
"[| Transset(i); !!x. x:i ==> Transset(x) |] ==> Ord(i)"; |
435 | 93 |
by (REPEAT (ares_tac (prems@[ballI,conjI]) 1)); |
760 | 94 |
qed "OrdI"; |
435 | 95 |
|
5321 | 96 |
Goalw [Ord_def] "Ord(i) ==> Transset(i)"; |
97 |
by (Blast_tac 1); |
|
760 | 98 |
qed "Ord_is_Transset"; |
435 | 99 |
|
5321 | 100 |
Goalw [Ord_def] |
435 | 101 |
"[| Ord(i); j:i |] ==> Transset(j) "; |
5321 | 102 |
by (Blast_tac 1); |
760 | 103 |
qed "Ord_contains_Transset"; |
435 | 104 |
|
105 |
(*** Lemmas for ordinals ***) |
|
106 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5137
diff
changeset
|
107 |
Goalw [Ord_def,Transset_def] "[| Ord(i); j:i |] ==> Ord(j)"; |
2925 | 108 |
by (Blast_tac 1); |
760 | 109 |
qed "Ord_in_Ord"; |
435 | 110 |
|
111 |
(* Ord(succ(j)) ==> Ord(j) *) |
|
9907 | 112 |
bind_thm ("Ord_succD", succI1 RSN (2, Ord_in_Ord)); |
435 | 113 |
|
3016 | 114 |
AddSDs [Ord_succD]; |
115 |
||
5137 | 116 |
Goal "[| Ord(i); Transset(j); j<=i |] ==> Ord(j)"; |
435 | 117 |
by (REPEAT (ares_tac [OrdI] 1 |
118 |
ORELSE eresolve_tac [Ord_contains_Transset, subsetD] 1)); |
|
760 | 119 |
qed "Ord_subset_Ord"; |
435 | 120 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5137
diff
changeset
|
121 |
Goalw [Ord_def,Transset_def] "[| j:i; Ord(i) |] ==> j<=i"; |
2925 | 122 |
by (Blast_tac 1); |
760 | 123 |
qed "OrdmemD"; |
435 | 124 |
|
5137 | 125 |
Goal "[| i:j; j:k; Ord(k) |] ==> i:k"; |
435 | 126 |
by (REPEAT (ares_tac [OrdmemD RS subsetD] 1)); |
760 | 127 |
qed "Ord_trans"; |
435 | 128 |
|
5137 | 129 |
Goal "[| i:j; Ord(j) |] ==> succ(i) <= j"; |
435 | 130 |
by (REPEAT (ares_tac [OrdmemD RSN (2,succ_subsetI)] 1)); |
760 | 131 |
qed "Ord_succ_subsetI"; |
435 | 132 |
|
133 |
||
134 |
(*** The construction of ordinals: 0, succ, Union ***) |
|
135 |
||
5067 | 136 |
Goal "Ord(0)"; |
435 | 137 |
by (REPEAT (ares_tac [OrdI,Transset_0] 1 ORELSE etac emptyE 1)); |
760 | 138 |
qed "Ord_0"; |
435 | 139 |
|
5137 | 140 |
Goal "Ord(i) ==> Ord(succ(i))"; |
435 | 141 |
by (REPEAT (ares_tac [OrdI,Transset_succ] 1 |
142 |
ORELSE eresolve_tac [succE,ssubst,Ord_is_Transset, |
|
1461 | 143 |
Ord_contains_Transset] 1)); |
760 | 144 |
qed "Ord_succ"; |
435 | 145 |
|
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
146 |
bind_thm ("Ord_1", Ord_0 RS Ord_succ); |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
147 |
|
5067 | 148 |
Goal "Ord(succ(i)) <-> Ord(i)"; |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6176
diff
changeset
|
149 |
by (blast_tac (claset() addIs [Ord_succ]) 1); |
760 | 150 |
qed "Ord_succ_iff"; |
435 | 151 |
|
2469 | 152 |
Addsimps [Ord_0, Ord_succ_iff]; |
153 |
AddSIs [Ord_0, Ord_succ]; |
|
6153 | 154 |
AddTCs [Ord_0, Ord_succ]; |
2469 | 155 |
|
5137 | 156 |
Goalw [Ord_def] "[| Ord(i); Ord(j) |] ==> Ord(i Un j)"; |
4091 | 157 |
by (blast_tac (claset() addSIs [Transset_Un]) 1); |
760 | 158 |
qed "Ord_Un"; |
435 | 159 |
|
5137 | 160 |
Goalw [Ord_def] "[| Ord(i); Ord(j) |] ==> Ord(i Int j)"; |
4091 | 161 |
by (blast_tac (claset() addSIs [Transset_Int]) 1); |
760 | 162 |
qed "Ord_Int"; |
6153 | 163 |
AddTCs [Ord_Un, Ord_Int]; |
435 | 164 |
|
5321 | 165 |
val nonempty::prems = Goal |
435 | 166 |
"[| j:A; !!i. i:A ==> Ord(i) |] ==> Ord(Inter(A))"; |
167 |
by (rtac (nonempty RS Transset_Inter_family RS OrdI) 1); |
|
168 |
by (rtac Ord_is_Transset 1); |
|
169 |
by (REPEAT (ares_tac ([Ord_contains_Transset,nonempty]@prems) 1 |
|
170 |
ORELSE etac InterD 1)); |
|
760 | 171 |
qed "Ord_Inter"; |
435 | 172 |
|
5321 | 173 |
val jmemA::prems = Goal |
435 | 174 |
"[| j:A; !!x. x:A ==> Ord(B(x)) |] ==> Ord(INT x:A. B(x))"; |
175 |
by (rtac (jmemA RS RepFunI RS Ord_Inter) 1); |
|
176 |
by (etac RepFunE 1); |
|
177 |
by (etac ssubst 1); |
|
178 |
by (eresolve_tac prems 1); |
|
760 | 179 |
qed "Ord_INT"; |
435 | 180 |
|
181 |
(*There is no set of all ordinals, for then it would contain itself*) |
|
5067 | 182 |
Goal "~ (ALL i. i:X <-> Ord(i))"; |
435 | 183 |
by (rtac notI 1); |
184 |
by (forw_inst_tac [("x", "X")] spec 1); |
|
4091 | 185 |
by (safe_tac (claset() addSEs [mem_irrefl])); |
435 | 186 |
by (swap_res_tac [Ord_is_Transset RSN (2,OrdI)] 1); |
2925 | 187 |
by (Blast_tac 2); |
437 | 188 |
by (rewtac Transset_def); |
4152 | 189 |
by Safe_tac; |
2469 | 190 |
by (Asm_full_simp_tac 1); |
435 | 191 |
by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 1)); |
760 | 192 |
qed "ON_class"; |
435 | 193 |
|
194 |
(*** < is 'less than' for ordinals ***) |
|
195 |
||
5137 | 196 |
Goalw [lt_def] "[| i:j; Ord(j) |] ==> i<j"; |
435 | 197 |
by (REPEAT (ares_tac [conjI] 1)); |
760 | 198 |
qed "ltI"; |
435 | 199 |
|
5321 | 200 |
val major::prems = Goalw [lt_def] |
435 | 201 |
"[| i<j; [| i:j; Ord(i); Ord(j) |] ==> P |] ==> P"; |
202 |
by (rtac (major RS conjE) 1); |
|
203 |
by (REPEAT (ares_tac (prems@[Ord_in_Ord]) 1)); |
|
760 | 204 |
qed "ltE"; |
435 | 205 |
|
5137 | 206 |
Goal "i<j ==> i:j"; |
435 | 207 |
by (etac ltE 1); |
208 |
by (assume_tac 1); |
|
760 | 209 |
qed "ltD"; |
435 | 210 |
|
5067 | 211 |
Goalw [lt_def] "~ i<0"; |
2925 | 212 |
by (Blast_tac 1); |
760 | 213 |
qed "not_lt0"; |
435 | 214 |
|
2469 | 215 |
Addsimps [not_lt0]; |
216 |
||
5137 | 217 |
Goal "j<i ==> Ord(j)"; |
1461 | 218 |
by (etac ltE 1 THEN assume_tac 1); |
830
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
lcp
parents:
782
diff
changeset
|
219 |
qed "lt_Ord"; |
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
lcp
parents:
782
diff
changeset
|
220 |
|
5137 | 221 |
Goal "j<i ==> Ord(i)"; |
1461 | 222 |
by (etac ltE 1 THEN assume_tac 1); |
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
223 |
qed "lt_Ord2"; |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
224 |
|
1034 | 225 |
(* "ja le j ==> Ord(j)" *) |
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
226 |
bind_thm ("le_Ord2", lt_Ord2 RS Ord_succD); |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
227 |
|
435 | 228 |
(* i<0 ==> R *) |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
772
diff
changeset
|
229 |
bind_thm ("lt0E", not_lt0 RS notE); |
435 | 230 |
|
5137 | 231 |
Goal "[| i<j; j<k |] ==> i<k"; |
4091 | 232 |
by (blast_tac (claset() addSIs [ltI] addSEs [ltE] addIs [Ord_trans]) 1); |
760 | 233 |
qed "lt_trans"; |
435 | 234 |
|
5465 | 235 |
Goalw [lt_def] "i<j ==> ~ (j<i)"; |
236 |
by (blast_tac (claset() addEs [mem_asym]) 1); |
|
237 |
qed "lt_not_sym"; |
|
238 |
||
239 |
(* [| i<j; ~P ==> j<i |] ==> P *) |
|
240 |
bind_thm ("lt_asym", lt_not_sym RS swap); |
|
435 | 241 |
|
9907 | 242 |
val [major]= goal (the_context ()) "i<i ==> P"; |
9180 | 243 |
by (rtac (major RS (major RS lt_asym)) 1) ; |
244 |
qed "lt_irrefl"; |
|
435 | 245 |
|
9180 | 246 |
Goal "~ i<i"; |
247 |
by (rtac notI 1); |
|
248 |
by (etac lt_irrefl 1) ; |
|
249 |
qed "lt_not_refl"; |
|
435 | 250 |
|
2469 | 251 |
AddSEs [lt_irrefl, lt0E]; |
252 |
||
435 | 253 |
(** le is less than or equals; recall i le j abbrevs i<succ(j) !! **) |
254 |
||
5067 | 255 |
Goalw [lt_def] "i le j <-> i<j | (i=j & Ord(j))"; |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6176
diff
changeset
|
256 |
by (Blast_tac 1); |
760 | 257 |
qed "le_iff"; |
435 | 258 |
|
772 | 259 |
(*Equivalently, i<j ==> i < succ(j)*) |
5137 | 260 |
Goal "i<j ==> i le j"; |
4091 | 261 |
by (asm_simp_tac (simpset() addsimps [le_iff]) 1); |
760 | 262 |
qed "leI"; |
435 | 263 |
|
5137 | 264 |
Goal "[| i=j; Ord(j) |] ==> i le j"; |
4091 | 265 |
by (asm_simp_tac (simpset() addsimps [le_iff]) 1); |
760 | 266 |
qed "le_eqI"; |
435 | 267 |
|
9907 | 268 |
bind_thm ("le_refl", refl RS le_eqI); |
435 | 269 |
|
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6176
diff
changeset
|
270 |
Goal "i le i <-> Ord(i)"; |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6176
diff
changeset
|
271 |
by (asm_simp_tac (simpset() addsimps [lt_not_refl, le_iff]) 1); |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6176
diff
changeset
|
272 |
qed "le_refl_iff"; |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6176
diff
changeset
|
273 |
|
9302 | 274 |
AddIffs [le_refl_iff]; |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6176
diff
changeset
|
275 |
|
5321 | 276 |
val [prem] = Goal "(~ (i=j & Ord(j)) ==> i<j) ==> i le j"; |
435 | 277 |
by (rtac (disjCI RS (le_iff RS iffD2)) 1); |
278 |
by (etac prem 1); |
|
760 | 279 |
qed "leCI"; |
435 | 280 |
|
5321 | 281 |
val major::prems = Goal |
435 | 282 |
"[| i le j; i<j ==> P; [| i=j; Ord(j) |] ==> P |] ==> P"; |
283 |
by (rtac (major RS (le_iff RS iffD1 RS disjE)) 1); |
|
284 |
by (DEPTH_SOLVE (ares_tac prems 1 ORELSE etac conjE 1)); |
|
760 | 285 |
qed "leE"; |
435 | 286 |
|
5137 | 287 |
Goal "[| i le j; j le i |] ==> i=j"; |
4091 | 288 |
by (asm_full_simp_tac (simpset() addsimps [le_iff]) 1); |
289 |
by (blast_tac (claset() addEs [lt_asym]) 1); |
|
760 | 290 |
qed "le_anti_sym"; |
435 | 291 |
|
5067 | 292 |
Goal "i le 0 <-> i=0"; |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6176
diff
changeset
|
293 |
by (blast_tac (claset() addSEs [leE]) 1); |
760 | 294 |
qed "le0_iff"; |
435 | 295 |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
772
diff
changeset
|
296 |
bind_thm ("le0D", le0_iff RS iffD1); |
435 | 297 |
|
2469 | 298 |
AddSDs [le0D]; |
299 |
Addsimps [le0_iff]; |
|
300 |
||
4091 | 301 |
val le_cs = claset() addSIs [leCI] addSEs [leE] addEs [lt_asym]; |
435 | 302 |
|
303 |
||
304 |
(*** Natural Deduction rules for Memrel ***) |
|
305 |
||
5067 | 306 |
Goalw [Memrel_def] "<a,b> : Memrel(A) <-> a:b & a:A & b:A"; |
2925 | 307 |
by (Blast_tac 1); |
760 | 308 |
qed "Memrel_iff"; |
9842 | 309 |
Addsimps [Memrel_iff]; |
310 |
(*MemrelI/E give better speed than AddIffs here*) |
|
435 | 311 |
|
5137 | 312 |
Goal "[| a: b; a: A; b: A |] ==> <a,b> : Memrel(A)"; |
9842 | 313 |
by Auto_tac; |
760 | 314 |
qed "MemrelI"; |
435 | 315 |
|
5321 | 316 |
val [major,minor] = Goal |
435 | 317 |
"[| <a,b> : Memrel(A); \ |
318 |
\ [| a: A; b: A; a:b |] ==> P \ |
|
319 |
\ |] ==> P"; |
|
320 |
by (rtac (major RS (Memrel_iff RS iffD1) RS conjE) 1); |
|
321 |
by (etac conjE 1); |
|
322 |
by (rtac minor 1); |
|
323 |
by (REPEAT (assume_tac 1)); |
|
760 | 324 |
qed "MemrelE"; |
435 | 325 |
|
2925 | 326 |
AddSIs [MemrelI]; |
327 |
AddSEs [MemrelE]; |
|
328 |
||
5067 | 329 |
Goalw [Memrel_def] "Memrel(A) <= A*A"; |
2925 | 330 |
by (Blast_tac 1); |
830
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
lcp
parents:
782
diff
changeset
|
331 |
qed "Memrel_type"; |
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
lcp
parents:
782
diff
changeset
|
332 |
|
5137 | 333 |
Goalw [Memrel_def] "A<=B ==> Memrel(A) <= Memrel(B)"; |
2925 | 334 |
by (Blast_tac 1); |
830
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
lcp
parents:
782
diff
changeset
|
335 |
qed "Memrel_mono"; |
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
lcp
parents:
782
diff
changeset
|
336 |
|
5067 | 337 |
Goalw [Memrel_def] "Memrel(0) = 0"; |
2925 | 338 |
by (Blast_tac 1); |
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
339 |
qed "Memrel_0"; |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
340 |
|
5067 | 341 |
Goalw [Memrel_def] "Memrel(1) = 0"; |
2925 | 342 |
by (Blast_tac 1); |
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
343 |
qed "Memrel_1"; |
830
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
lcp
parents:
782
diff
changeset
|
344 |
|
2469 | 345 |
Addsimps [Memrel_0, Memrel_1]; |
346 |
||
435 | 347 |
(*The membership relation (as a set) is well-founded. |
348 |
Proof idea: show A<=B by applying the foundation axiom to A-B *) |
|
5067 | 349 |
Goalw [wf_def] "wf(Memrel(A))"; |
435 | 350 |
by (EVERY1 [rtac (foundation RS disjE RS allI), |
1461 | 351 |
etac disjI1, |
352 |
etac bexE, |
|
353 |
rtac (impI RS allI RS bexI RS disjI2), |
|
354 |
etac MemrelE, |
|
355 |
etac bspec, |
|
356 |
REPEAT o assume_tac]); |
|
760 | 357 |
qed "wf_Memrel"; |
435 | 358 |
|
359 |
(*Transset(i) does not suffice, though ALL j:i.Transset(j) does*) |
|
5067 | 360 |
Goalw [Ord_def, Transset_def, trans_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
361 |
"Ord(i) ==> trans(Memrel(i))"; |
2925 | 362 |
by (Blast_tac 1); |
760 | 363 |
qed "trans_Memrel"; |
435 | 364 |
|
365 |
(*If Transset(A) then Memrel(A) internalizes the membership relation below A*) |
|
5067 | 366 |
Goalw [Transset_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
367 |
"Transset(A) ==> <a,b> : Memrel(A) <-> a:b & b:A"; |
2925 | 368 |
by (Blast_tac 1); |
760 | 369 |
qed "Transset_Memrel_iff"; |
435 | 370 |
|
371 |
||
372 |
(*** Transfinite induction ***) |
|
373 |
||
374 |
(*Epsilon induction over a transitive set*) |
|
5321 | 375 |
val major::prems = Goalw [Transset_def] |
435 | 376 |
"[| i: k; Transset(k); \ |
377 |
\ !!x.[| x: k; ALL y:x. P(y) |] ==> P(x) \ |
|
378 |
\ |] ==> P(i)"; |
|
379 |
by (rtac (major RS (wf_Memrel RS wf_induct2)) 1); |
|
2925 | 380 |
by (Blast_tac 1); |
435 | 381 |
by (resolve_tac prems 1); |
382 |
by (assume_tac 1); |
|
383 |
by (cut_facts_tac prems 1); |
|
2925 | 384 |
by (Blast_tac 1); |
760 | 385 |
qed "Transset_induct"; |
435 | 386 |
|
387 |
(*Induction over an ordinal*) |
|
9907 | 388 |
bind_thm ("Ord_induct", Ord_is_Transset RSN (2, Transset_induct)); |
435 | 389 |
|
390 |
(*Induction over the class of ordinals -- a useful corollary of Ord_induct*) |
|
5321 | 391 |
val [major,indhyp] = Goal |
435 | 392 |
"[| Ord(i); \ |
393 |
\ !!x.[| Ord(x); ALL y:x. P(y) |] ==> P(x) \ |
|
394 |
\ |] ==> P(i)"; |
|
395 |
by (rtac (major RS Ord_succ RS (succI1 RS Ord_induct)) 1); |
|
396 |
by (rtac indhyp 1); |
|
397 |
by (rtac (major RS Ord_succ RS Ord_in_Ord) 1); |
|
398 |
by (REPEAT (assume_tac 1)); |
|
760 | 399 |
qed "trans_induct"; |
435 | 400 |
|
401 |
(*Perform induction on i, then prove the Ord(i) subgoal using prems. *) |
|
402 |
fun trans_ind_tac a prems i = |
|
403 |
EVERY [res_inst_tac [("i",a)] trans_induct i, |
|
1461 | 404 |
rename_last_tac a ["1"] (i+1), |
405 |
ares_tac prems i]; |
|
435 | 406 |
|
407 |
||
408 |
(*** Fundamental properties of the epsilon ordering (< on ordinals) ***) |
|
409 |
||
410 |
(*Finds contradictions for the following proof*) |
|
411 |
val Ord_trans_tac = EVERY' [etac notE, etac Ord_trans, REPEAT o atac]; |
|
412 |
||
413 |
(** Proving that < is a linear ordering on the ordinals **) |
|
414 |
||
5321 | 415 |
Goal "Ord(i) ==> (ALL j. Ord(j) --> i:j | i=j | j:i)"; |
416 |
by (etac trans_induct 1); |
|
435 | 417 |
by (rtac (impI RS allI) 1); |
418 |
by (trans_ind_tac "j" [] 1); |
|
2493 | 419 |
by (DEPTH_SOLVE (Step_tac 1 ORELSE Ord_trans_tac 1)); |
3736
39ee3d31cfbc
Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents:
3016
diff
changeset
|
420 |
qed_spec_mp "Ord_linear"; |
435 | 421 |
|
422 |
(*The trichotomy law for ordinals!*) |
|
5321 | 423 |
val ordi::ordj::prems = Goalw [lt_def] |
435 | 424 |
"[| Ord(i); Ord(j); i<j ==> P; i=j ==> P; j<i ==> P |] ==> P"; |
425 |
by (rtac ([ordi,ordj] MRS Ord_linear RS disjE) 1); |
|
426 |
by (etac disjE 2); |
|
427 |
by (DEPTH_SOLVE (ares_tac ([ordi,ordj,conjI] @ prems) 1)); |
|
760 | 428 |
qed "Ord_linear_lt"; |
435 | 429 |
|
5321 | 430 |
val prems = Goal |
435 | 431 |
"[| Ord(i); Ord(j); i<j ==> P; j le i ==> P |] ==> P"; |
432 |
by (res_inst_tac [("i","i"),("j","j")] Ord_linear_lt 1); |
|
433 |
by (DEPTH_SOLVE (ares_tac ([leI, sym RS le_eqI] @ prems) 1)); |
|
760 | 434 |
qed "Ord_linear2"; |
435 | 435 |
|
5321 | 436 |
val prems = Goal |
435 | 437 |
"[| Ord(i); Ord(j); i le j ==> P; j le i ==> P |] ==> P"; |
438 |
by (res_inst_tac [("i","i"),("j","j")] Ord_linear_lt 1); |
|
439 |
by (DEPTH_SOLVE (ares_tac ([leI,le_eqI] @ prems) 1)); |
|
760 | 440 |
qed "Ord_linear_le"; |
435 | 441 |
|
5137 | 442 |
Goal "j le i ==> ~ i<j"; |
2925 | 443 |
by (blast_tac le_cs 1); |
760 | 444 |
qed "le_imp_not_lt"; |
435 | 445 |
|
5137 | 446 |
Goal "[| ~ i<j; Ord(i); Ord(j) |] ==> j le i"; |
435 | 447 |
by (res_inst_tac [("i","i"),("j","j")] Ord_linear2 1); |
448 |
by (REPEAT (SOMEGOAL assume_tac)); |
|
2925 | 449 |
by (blast_tac le_cs 1); |
760 | 450 |
qed "not_lt_imp_le"; |
435 | 451 |
|
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
452 |
(** Some rewrite rules for <, le **) |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
453 |
|
5137 | 454 |
Goalw [lt_def] "Ord(j) ==> i:j <-> i<j"; |
2925 | 455 |
by (Blast_tac 1); |
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
456 |
qed "Ord_mem_iff_lt"; |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
457 |
|
5137 | 458 |
Goal "[| Ord(i); Ord(j) |] ==> ~ i<j <-> j le i"; |
435 | 459 |
by (REPEAT (ares_tac [iffI, le_imp_not_lt, not_lt_imp_le] 1)); |
760 | 460 |
qed "not_lt_iff_le"; |
435 | 461 |
|
5137 | 462 |
Goal "[| Ord(i); Ord(j) |] ==> ~ i le j <-> j<i"; |
4091 | 463 |
by (asm_simp_tac (simpset() addsimps [not_lt_iff_le RS iff_sym]) 1); |
760 | 464 |
qed "not_le_iff_lt"; |
435 | 465 |
|
1610 | 466 |
(*This is identical to 0<succ(i) *) |
5137 | 467 |
Goal "Ord(i) ==> 0 le i"; |
435 | 468 |
by (etac (not_lt_iff_le RS iffD1) 1); |
469 |
by (REPEAT (resolve_tac [Ord_0, not_lt0] 1)); |
|
760 | 470 |
qed "Ord_0_le"; |
435 | 471 |
|
5137 | 472 |
Goal "[| Ord(i); i~=0 |] ==> 0<i"; |
435 | 473 |
by (etac (not_le_iff_lt RS iffD1) 1); |
474 |
by (rtac Ord_0 1); |
|
2925 | 475 |
by (Blast_tac 1); |
760 | 476 |
qed "Ord_0_lt"; |
435 | 477 |
|
9492
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9302
diff
changeset
|
478 |
Goal "Ord(i) ==> i~=0 <-> 0<i"; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9302
diff
changeset
|
479 |
by (blast_tac (claset() addIs [Ord_0_lt]) 1); |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9302
diff
changeset
|
480 |
qed "Ord_0_lt_iff"; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9302
diff
changeset
|
481 |
|
9872 | 482 |
|
435 | 483 |
(*** Results about less-than or equals ***) |
484 |
||
485 |
(** For ordinals, j<=i (subset) implies j le i (less-than or equals) **) |
|
486 |
||
9872 | 487 |
Goal "0 le succ(x) <-> Ord(x)"; |
488 |
by (blast_tac (claset() addIs [Ord_0_le] addEs [ltE]) 1); |
|
489 |
qed "zero_le_succ_iff"; |
|
490 |
AddIffs [zero_le_succ_iff]; |
|
491 |
||
5137 | 492 |
Goal "[| j<=i; Ord(i); Ord(j) |] ==> j le i"; |
435 | 493 |
by (rtac (not_lt_iff_le RS iffD1) 1); |
494 |
by (assume_tac 1); |
|
495 |
by (assume_tac 1); |
|
4091 | 496 |
by (blast_tac (claset() addEs [ltE, mem_irrefl]) 1); |
760 | 497 |
qed "subset_imp_le"; |
435 | 498 |
|
5137 | 499 |
Goal "i le j ==> i<=j"; |
435 | 500 |
by (etac leE 1); |
2925 | 501 |
by (Blast_tac 2); |
502 |
by (blast_tac (subset_cs addIs [OrdmemD] addEs [ltE]) 1); |
|
760 | 503 |
qed "le_imp_subset"; |
435 | 504 |
|
5067 | 505 |
Goal "j le i <-> j<=i & Ord(i) & Ord(j)"; |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6176
diff
changeset
|
506 |
by (blast_tac (claset() addDs [subset_imp_le, le_imp_subset] addEs [ltE]) 1); |
760 | 507 |
qed "le_subset_iff"; |
435 | 508 |
|
5067 | 509 |
Goal "i le succ(j) <-> i le j | i=succ(j) & Ord(i)"; |
4091 | 510 |
by (simp_tac (simpset() addsimps [le_iff]) 1); |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6176
diff
changeset
|
511 |
by (Blast_tac 1); |
760 | 512 |
qed "le_succ_iff"; |
435 | 513 |
|
514 |
(*Just a variant of subset_imp_le*) |
|
5321 | 515 |
val [ordi,ordj,minor] = Goal |
435 | 516 |
"[| Ord(i); Ord(j); !!x. x<j ==> x<i |] ==> j le i"; |
517 |
by (REPEAT_FIRST (ares_tac [notI RS not_lt_imp_le, ordi, ordj])); |
|
437 | 518 |
by (etac (minor RS lt_irrefl) 1); |
760 | 519 |
qed "all_lt_imp_le"; |
435 | 520 |
|
521 |
(** Transitive laws **) |
|
522 |
||
5137 | 523 |
Goal "[| i le j; j<k |] ==> i<k"; |
4091 | 524 |
by (blast_tac (claset() addSEs [leE] addIs [lt_trans]) 1); |
760 | 525 |
qed "lt_trans1"; |
435 | 526 |
|
5137 | 527 |
Goal "[| i<j; j le k |] ==> i<k"; |
4091 | 528 |
by (blast_tac (claset() addSEs [leE] addIs [lt_trans]) 1); |
760 | 529 |
qed "lt_trans2"; |
435 | 530 |
|
5137 | 531 |
Goal "[| i le j; j le k |] ==> i le k"; |
435 | 532 |
by (REPEAT (ares_tac [lt_trans1] 1)); |
760 | 533 |
qed "le_trans"; |
435 | 534 |
|
5137 | 535 |
Goal "i<j ==> succ(i) le j"; |
435 | 536 |
by (rtac (not_lt_iff_le RS iffD1) 1); |
2925 | 537 |
by (blast_tac le_cs 3); |
4091 | 538 |
by (ALLGOALS (blast_tac (claset() addEs [ltE]))); |
760 | 539 |
qed "succ_leI"; |
435 | 540 |
|
830
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
lcp
parents:
782
diff
changeset
|
541 |
(*Identical to succ(i) < succ(j) ==> i<j *) |
5137 | 542 |
Goal "succ(i) le j ==> i<j"; |
435 | 543 |
by (rtac (not_le_iff_lt RS iffD1) 1); |
4475 | 544 |
by (blast_tac le_cs 3); |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6176
diff
changeset
|
545 |
by (ALLGOALS (blast_tac (claset() addEs [ltE]))); |
760 | 546 |
qed "succ_leE"; |
435 | 547 |
|
5067 | 548 |
Goal "succ(i) le j <-> i<j"; |
435 | 549 |
by (REPEAT (ares_tac [iffI,succ_leI,succ_leE] 1)); |
760 | 550 |
qed "succ_le_iff"; |
435 | 551 |
|
2469 | 552 |
Addsimps [succ_le_iff]; |
553 |
||
5137 | 554 |
Goal "succ(i) le succ(j) ==> i le j"; |
4091 | 555 |
by (blast_tac (claset() addSDs [succ_leE]) 1); |
830
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
lcp
parents:
782
diff
changeset
|
556 |
qed "succ_le_imp_le"; |
18240b5d8a06
Moved Transset_includes_summands and Transset_sum_Int_subset to
lcp
parents:
782
diff
changeset
|
557 |
|
6176
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
558 |
Goal "[| i <= j; j<k; Ord(i) |] ==> i<k"; |
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
559 |
by (resolve_tac [subset_imp_le RS lt_trans1] 1); |
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
560 |
by (REPEAT (eresolve_tac [asm_rl, ltE] 1)); |
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
561 |
qed "lt_subset_trans"; |
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
562 |
|
435 | 563 |
(** Union and Intersection **) |
564 |
||
5137 | 565 |
Goal "[| Ord(i); Ord(j) |] ==> i le i Un j"; |
435 | 566 |
by (rtac (Un_upper1 RS subset_imp_le) 1); |
567 |
by (REPEAT (ares_tac [Ord_Un] 1)); |
|
760 | 568 |
qed "Un_upper1_le"; |
435 | 569 |
|
5137 | 570 |
Goal "[| Ord(i); Ord(j) |] ==> j le i Un j"; |
435 | 571 |
by (rtac (Un_upper2 RS subset_imp_le) 1); |
572 |
by (REPEAT (ares_tac [Ord_Un] 1)); |
|
760 | 573 |
qed "Un_upper2_le"; |
435 | 574 |
|
575 |
(*Replacing k by succ(k') yields the similar rule for le!*) |
|
5137 | 576 |
Goal "[| i<k; j<k |] ==> i Un j < k"; |
435 | 577 |
by (res_inst_tac [("i","i"),("j","j")] Ord_linear_le 1); |
2033 | 578 |
by (stac Un_commute 4); |
4091 | 579 |
by (asm_full_simp_tac (simpset() addsimps [le_subset_iff, subset_Un_iff]) 4); |
580 |
by (asm_full_simp_tac (simpset() addsimps [le_subset_iff, subset_Un_iff]) 3); |
|
435 | 581 |
by (REPEAT (eresolve_tac [asm_rl, ltE] 1)); |
760 | 582 |
qed "Un_least_lt"; |
435 | 583 |
|
5137 | 584 |
Goal "[| Ord(i); Ord(j) |] ==> i Un j < k <-> i<k & j<k"; |
4091 | 585 |
by (safe_tac (claset() addSIs [Un_least_lt])); |
437 | 586 |
by (rtac (Un_upper2_le RS lt_trans1) 2); |
587 |
by (rtac (Un_upper1_le RS lt_trans1) 1); |
|
435 | 588 |
by (REPEAT_SOME assume_tac); |
760 | 589 |
qed "Un_least_lt_iff"; |
435 | 590 |
|
9907 | 591 |
val [ordi,ordj,ordk] = goal (the_context ()) |
435 | 592 |
"[| Ord(i); Ord(j); Ord(k) |] ==> i Un j : k <-> i:k & j:k"; |
8551 | 593 |
by (cut_inst_tac [("k","k")] ([ordi,ordj] MRS Un_least_lt_iff) 1); |
4091 | 594 |
by (asm_full_simp_tac (simpset() addsimps [lt_def,ordi,ordj,ordk]) 1); |
760 | 595 |
qed "Un_least_mem_iff"; |
435 | 596 |
|
597 |
(*Replacing k by succ(k') yields the similar rule for le!*) |
|
5137 | 598 |
Goal "[| i<k; j<k |] ==> i Int j < k"; |
435 | 599 |
by (res_inst_tac [("i","i"),("j","j")] Ord_linear_le 1); |
2033 | 600 |
by (stac Int_commute 4); |
4091 | 601 |
by (asm_full_simp_tac (simpset() addsimps [le_subset_iff, subset_Int_iff]) 4); |
602 |
by (asm_full_simp_tac (simpset() addsimps [le_subset_iff, subset_Int_iff]) 3); |
|
435 | 603 |
by (REPEAT (eresolve_tac [asm_rl, ltE] 1)); |
760 | 604 |
qed "Int_greatest_lt"; |
435 | 605 |
|
606 |
(*FIXME: the Intersection duals are missing!*) |
|
607 |
||
608 |
||
609 |
(*** Results about limits ***) |
|
610 |
||
5321 | 611 |
val prems = Goal "[| !!i. i:A ==> Ord(i) |] ==> Ord(Union(A))"; |
435 | 612 |
by (rtac (Ord_is_Transset RS Transset_Union_family RS OrdI) 1); |
613 |
by (REPEAT (etac UnionE 1 ORELSE ares_tac ([Ord_contains_Transset]@prems) 1)); |
|
760 | 614 |
qed "Ord_Union"; |
435 | 615 |
|
5321 | 616 |
val prems = Goal |
435 | 617 |
"[| !!x. x:A ==> Ord(B(x)) |] ==> Ord(UN x:A. B(x))"; |
618 |
by (rtac Ord_Union 1); |
|
619 |
by (etac RepFunE 1); |
|
620 |
by (etac ssubst 1); |
|
621 |
by (eresolve_tac prems 1); |
|
760 | 622 |
qed "Ord_UN"; |
435 | 623 |
|
624 |
(* No < version; consider (UN i:nat.i)=nat *) |
|
5321 | 625 |
val [ordi,limit] = Goal |
435 | 626 |
"[| Ord(i); !!x. x:A ==> b(x) le i |] ==> (UN x:A. b(x)) le i"; |
627 |
by (rtac (le_imp_subset RS UN_least RS subset_imp_le) 1); |
|
628 |
by (REPEAT (ares_tac [ordi, Ord_UN, limit] 1 ORELSE etac (limit RS ltE) 1)); |
|
760 | 629 |
qed "UN_least_le"; |
435 | 630 |
|
5321 | 631 |
val [jlti,limit] = Goal |
435 | 632 |
"[| j<i; !!x. x:A ==> b(x)<j |] ==> (UN x:A. succ(b(x))) < i"; |
633 |
by (rtac (jlti RS ltE) 1); |
|
634 |
by (rtac (UN_least_le RS lt_trans2) 1); |
|
635 |
by (REPEAT (ares_tac [jlti, succ_leI, limit] 1)); |
|
760 | 636 |
qed "UN_succ_least_lt"; |
435 | 637 |
|
12485 | 638 |
Goal "[| a: A; i le b(a); Ord(UN x:A. b(x)) |] ==> i le (UN x:A. b(x))"; |
639 |
by (ftac ltD 1); |
|
435 | 640 |
by (rtac (le_imp_subset RS subset_trans RS subset_imp_le) 1); |
12485 | 641 |
by (REPEAT (ares_tac [lt_Ord, UN_upper] 1)); |
760 | 642 |
qed "UN_upper_le"; |
435 | 643 |
|
5321 | 644 |
val [leprem] = Goal |
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
645 |
"[| !!x. x:A ==> c(x) le d(x) |] ==> (UN x:A. c(x)) le (UN x:A. d(x))"; |
1461 | 646 |
by (rtac UN_least_le 1); |
647 |
by (rtac UN_upper_le 2); |
|
12485 | 648 |
by (etac leprem 3); |
649 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [Ord_UN, leprem RS le_Ord2]))); |
|
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
650 |
qed "le_implies_UN_le_UN"; |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
651 |
|
5137 | 652 |
Goal "Ord(i) ==> (UN y:i. succ(y)) = i"; |
4091 | 653 |
by (blast_tac (claset() addIs [Ord_trans]) 1); |
760 | 654 |
qed "Ord_equality"; |
435 | 655 |
|
656 |
(*Holds for all transitive sets, not just ordinals*) |
|
5137 | 657 |
Goal "Ord(i) ==> Union(i) <= i"; |
4091 | 658 |
by (blast_tac (claset() addIs [Ord_trans]) 1); |
760 | 659 |
qed "Ord_Union_subset"; |
435 | 660 |
|
661 |
||
662 |
(*** Limit ordinals -- general properties ***) |
|
663 |
||
5137 | 664 |
Goalw [Limit_def] "Limit(i) ==> Union(i) = i"; |
4091 | 665 |
by (fast_tac (claset() addSIs [ltI] addSEs [ltE] addEs [Ord_trans]) 1); |
760 | 666 |
qed "Limit_Union_eq"; |
435 | 667 |
|
5137 | 668 |
Goalw [Limit_def] "Limit(i) ==> Ord(i)"; |
435 | 669 |
by (etac conjunct1 1); |
760 | 670 |
qed "Limit_is_Ord"; |
435 | 671 |
|
5137 | 672 |
Goalw [Limit_def] "Limit(i) ==> 0 < i"; |
435 | 673 |
by (etac (conjunct2 RS conjunct1) 1); |
760 | 674 |
qed "Limit_has_0"; |
435 | 675 |
|
5137 | 676 |
Goalw [Limit_def] "[| Limit(i); j<i |] ==> succ(j) < i"; |
2925 | 677 |
by (Blast_tac 1); |
760 | 678 |
qed "Limit_has_succ"; |
435 | 679 |
|
5067 | 680 |
Goalw [Limit_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
681 |
"[| 0<i; ALL y. succ(y) ~= i |] ==> Limit(i)"; |
435 | 682 |
by (safe_tac subset_cs); |
683 |
by (rtac (not_le_iff_lt RS iffD1) 2); |
|
2925 | 684 |
by (blast_tac le_cs 4); |
435 | 685 |
by (REPEAT (eresolve_tac [asm_rl, ltE, Ord_succ] 1)); |
760 | 686 |
qed "non_succ_LimitI"; |
435 | 687 |
|
5137 | 688 |
Goal "Limit(succ(i)) ==> P"; |
437 | 689 |
by (rtac lt_irrefl 1); |
690 |
by (rtac Limit_has_succ 1); |
|
691 |
by (assume_tac 1); |
|
692 |
by (etac (Limit_is_Ord RS Ord_succD RS le_refl) 1); |
|
760 | 693 |
qed "succ_LimitE"; |
12594 | 694 |
AddSEs [succ_LimitE]; |
695 |
||
696 |
Goal "~ Limit(succ(i))"; |
|
697 |
by (Blast_tac 1); |
|
698 |
qed "not_succ_Limit"; |
|
699 |
Addsimps [not_succ_Limit]; |
|
435 | 700 |
|
5137 | 701 |
Goal "[| Limit(i); i le succ(j) |] ==> i le j"; |
12594 | 702 |
by (blast_tac (claset() addSEs [leE]) 1); |
760 | 703 |
qed "Limit_le_succD"; |
435 | 704 |
|
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
705 |
(** Traditional 3-way case analysis on ordinals **) |
435 | 706 |
|
5137 | 707 |
Goal "Ord(i) ==> i=0 | (EX j. Ord(j) & i=succ(j)) | Limit(i)"; |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6176
diff
changeset
|
708 |
by (blast_tac (claset() addSIs [non_succ_LimitI, Ord_0_lt]) 1); |
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
709 |
qed "Ord_cases_disj"; |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
710 |
|
5321 | 711 |
val major::prems = Goal |
1461 | 712 |
"[| Ord(i); \ |
713 |
\ i=0 ==> P; \ |
|
714 |
\ !!j. [| Ord(j); i=succ(j) |] ==> P; \ |
|
715 |
\ Limit(i) ==> P \ |
|
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
716 |
\ |] ==> P"; |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
717 |
by (cut_facts_tac [major RS Ord_cases_disj] 1); |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
718 |
by (REPEAT (eresolve_tac (prems@[asm_rl, disjE, exE, conjE]) 1)); |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
719 |
qed "Ord_cases"; |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
720 |
|
5321 | 721 |
val major::prems = Goal |
1461 | 722 |
"[| Ord(i); \ |
723 |
\ P(0); \ |
|
724 |
\ !!x. [| Ord(x); P(x) |] ==> P(succ(x)); \ |
|
725 |
\ !!x. [| Limit(x); ALL y:x. P(y) |] ==> P(x) \ |
|
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
726 |
\ |] ==> P(i)"; |
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
727 |
by (resolve_tac [major RS trans_induct] 1); |
1461 | 728 |
by (etac Ord_cases 1); |
4091 | 729 |
by (ALLGOALS (blast_tac (claset() addIs prems))); |
851
f9172c4625f1
Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
lcp
parents:
830
diff
changeset
|
730 |
qed "trans_induct3"; |