author | clasohm |
Thu, 14 Mar 1996 12:19:49 +0100 | |
changeset 1577 | a84cc626ea69 |
parent 1461 | 6bcb44e4d6e5 |
child 1791 | 6b38717439c6 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/Order.ML |
435 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
435 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
789 | 6 |
Orders in Zermelo-Fraenkel Set Theory |
7 |
||
8 |
Results from the book "Set Theory: an Introduction to Independence Proofs" |
|
1461 | 9 |
by Ken Kunen. Chapter 1, section 6. |
435 | 10 |
*) |
11 |
||
12 |
open Order; |
|
13 |
||
14 |
val bij_apply_cs = ZF_cs addSEs [bij_converse_bij] |
|
15 |
addIs [bij_is_fun, apply_type]; |
|
16 |
||
17 |
(** Basic properties of the definitions **) |
|
18 |
||
789 | 19 |
(*needed?*) |
435 | 20 |
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def, asym_def] |
21 |
"!!r. part_ord(A,r) ==> asym(r Int A*A)"; |
|
22 |
by (fast_tac ZF_cs 1); |
|
760 | 23 |
qed "part_ord_Imp_asym"; |
435 | 24 |
|
789 | 25 |
val major::premx::premy::prems = goalw Order.thy [linear_def] |
26 |
"[| linear(A,r); x:A; y:A; \ |
|
27 |
\ <x,y>:r ==> P; x=y ==> P; <y,x>:r ==> P |] ==> P"; |
|
28 |
by (cut_facts_tac [major,premx,premy] 1); |
|
29 |
by (REPEAT_FIRST (eresolve_tac [ballE,disjE])); |
|
30 |
by (EVERY1 (map etac prems)); |
|
31 |
by (ALLGOALS contr_tac); |
|
32 |
qed "linearE"; |
|
33 |
||
34 |
(*Does the case analysis, deleting linear(A,r) and proving trivial subgoals*) |
|
35 |
val linear_case_tac = |
|
36 |
SELECT_GOAL (EVERY [etac linearE 1, assume_tac 1, assume_tac 1, |
|
1461 | 37 |
REPEAT_SOME (assume_tac ORELSE' contr_tac)]); |
789 | 38 |
|
39 |
(** General properties of well_ord **) |
|
40 |
||
41 |
goalw Order.thy [irrefl_def, part_ord_def, tot_ord_def, |
|
1461 | 42 |
trans_on_def, well_ord_def] |
789 | 43 |
"!!r. [| wf[A](r); linear(A,r) |] ==> well_ord(A,r)"; |
44 |
by (asm_simp_tac (ZF_ss addsimps [wf_on_not_refl]) 1); |
|
990
9ec3c7bd774e
Tried the new addss in many proofs, and tidied others involving simplification.
lcp
parents:
848
diff
changeset
|
45 |
by (fast_tac (ZF_cs addEs [linearE, wf_on_asym, wf_on_chain3]) 1); |
789 | 46 |
qed "well_ordI"; |
47 |
||
48 |
goalw Order.thy [well_ord_def] |
|
49 |
"!!r. well_ord(A,r) ==> wf[A](r)"; |
|
50 |
by (safe_tac ZF_cs); |
|
51 |
qed "well_ord_is_wf"; |
|
52 |
||
53 |
goalw Order.thy [well_ord_def, tot_ord_def, part_ord_def] |
|
54 |
"!!r. well_ord(A,r) ==> trans[A](r)"; |
|
55 |
by (safe_tac ZF_cs); |
|
56 |
qed "well_ord_is_trans_on"; |
|
484 | 57 |
|
789 | 58 |
goalw Order.thy [well_ord_def, tot_ord_def] |
59 |
"!!r. well_ord(A,r) ==> linear(A,r)"; |
|
60 |
by (fast_tac ZF_cs 1); |
|
61 |
qed "well_ord_is_linear"; |
|
62 |
||
63 |
||
64 |
(** Derived rules for pred(A,x,r) **) |
|
65 |
||
66 |
goalw Order.thy [pred_def] "y : pred(A,x,r) <-> <y,x>:r & y:A"; |
|
67 |
by (fast_tac ZF_cs 1); |
|
68 |
qed "pred_iff"; |
|
69 |
||
70 |
bind_thm ("predI", conjI RS (pred_iff RS iffD2)); |
|
71 |
||
72 |
val [major,minor] = goalw Order.thy [pred_def] |
|
73 |
"[| y: pred(A,x,r); [| y:A; <y,x>:r |] ==> P |] ==> P"; |
|
74 |
by (rtac (major RS CollectE) 1); |
|
75 |
by (REPEAT (ares_tac [minor] 1)); |
|
76 |
qed "predE"; |
|
77 |
||
78 |
goalw Order.thy [pred_def] "pred(A,x,r) <= r -`` {x}"; |
|
79 |
by (fast_tac ZF_cs 1); |
|
80 |
qed "pred_subset_under"; |
|
81 |
||
82 |
goalw Order.thy [pred_def] "pred(A,x,r) <= A"; |
|
83 |
by (fast_tac ZF_cs 1); |
|
84 |
qed "pred_subset"; |
|
85 |
||
86 |
goalw Order.thy [pred_def] |
|
87 |
"pred(pred(A,x,r), y, r) = pred(A,x,r) Int pred(A,y,r)"; |
|
88 |
by (fast_tac eq_cs 1); |
|
89 |
qed "pred_pred_eq"; |
|
90 |
||
91 |
goalw Order.thy [trans_on_def, pred_def] |
|
92 |
"!!r. [| trans[A](r); <y,x>:r; x:A; y:A \ |
|
93 |
\ |] ==> pred(pred(A,x,r), y, r) = pred(A,y,r)"; |
|
94 |
by (fast_tac eq_cs 1); |
|
95 |
qed "trans_pred_pred_eq"; |
|
96 |
||
836
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
97 |
|
789 | 98 |
(** The ordering's properties hold over all subsets of its domain |
99 |
[including initial segments of the form pred(A,x,r) **) |
|
484 | 100 |
|
101 |
(*Note: a relation s such that s<=r need not be a partial ordering*) |
|
102 |
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def] |
|
103 |
"!!A B r. [| part_ord(A,r); B<=A |] ==> part_ord(B,r)"; |
|
104 |
by (fast_tac ZF_cs 1); |
|
760 | 105 |
qed "part_ord_subset"; |
484 | 106 |
|
107 |
goalw Order.thy [linear_def] |
|
108 |
"!!A B r. [| linear(A,r); B<=A |] ==> linear(B,r)"; |
|
109 |
by (fast_tac ZF_cs 1); |
|
760 | 110 |
qed "linear_subset"; |
484 | 111 |
|
112 |
goalw Order.thy [tot_ord_def] |
|
113 |
"!!A B r. [| tot_ord(A,r); B<=A |] ==> tot_ord(B,r)"; |
|
114 |
by (fast_tac (ZF_cs addSEs [part_ord_subset, linear_subset]) 1); |
|
760 | 115 |
qed "tot_ord_subset"; |
484 | 116 |
|
117 |
goalw Order.thy [well_ord_def] |
|
118 |
"!!A B r. [| well_ord(A,r); B<=A |] ==> well_ord(B,r)"; |
|
119 |
by (fast_tac (ZF_cs addSEs [tot_ord_subset, wf_on_subset_A]) 1); |
|
760 | 120 |
qed "well_ord_subset"; |
484 | 121 |
|
122 |
||
836
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
123 |
(** Relations restricted to a smaller domain, by Krzysztof Grabczewski **) |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
124 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
125 |
goalw Order.thy [irrefl_def] "irrefl(A,r Int A*A) <-> irrefl(A,r)"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
126 |
by (fast_tac ZF_cs 1); |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
127 |
qed "irrefl_Int_iff"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
128 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
129 |
goalw Order.thy [trans_on_def] "trans[A](r Int A*A) <-> trans[A](r)"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
130 |
by (fast_tac ZF_cs 1); |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
131 |
qed "trans_on_Int_iff"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
132 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
133 |
goalw Order.thy [part_ord_def] "part_ord(A,r Int A*A) <-> part_ord(A,r)"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
134 |
by (simp_tac (ZF_ss addsimps [irrefl_Int_iff, trans_on_Int_iff]) 1); |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
135 |
qed "part_ord_Int_iff"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
136 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
137 |
goalw Order.thy [linear_def] "linear(A,r Int A*A) <-> linear(A,r)"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
138 |
by (fast_tac ZF_cs 1); |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
139 |
qed "linear_Int_iff"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
140 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
141 |
goalw Order.thy [tot_ord_def] "tot_ord(A,r Int A*A) <-> tot_ord(A,r)"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
142 |
by (simp_tac (ZF_ss addsimps [part_ord_Int_iff, linear_Int_iff]) 1); |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
143 |
qed "tot_ord_Int_iff"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
144 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
145 |
goalw Order.thy [wf_on_def, wf_def] "wf[A](r Int A*A) <-> wf[A](r)"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
146 |
by (fast_tac ZF_cs 1); |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
147 |
qed "wf_on_Int_iff"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
148 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
149 |
goalw Order.thy [well_ord_def] "well_ord(A,r Int A*A) <-> well_ord(A,r)"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
150 |
by (simp_tac (ZF_ss addsimps [tot_ord_Int_iff, wf_on_Int_iff]) 1); |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
151 |
qed "well_ord_Int_iff"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
152 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
153 |
|
848 | 154 |
(** Relations over the Empty Set **) |
155 |
||
156 |
goalw Order.thy [irrefl_def] "irrefl(0,r)"; |
|
157 |
by (fast_tac ZF_cs 1); |
|
158 |
qed "irrefl_0"; |
|
159 |
||
160 |
goalw Order.thy [trans_on_def] "trans[0](r)"; |
|
161 |
by (fast_tac ZF_cs 1); |
|
162 |
qed "trans_on_0"; |
|
163 |
||
164 |
goalw Order.thy [part_ord_def] "part_ord(0,r)"; |
|
165 |
by (simp_tac (ZF_ss addsimps [irrefl_0, trans_on_0]) 1); |
|
166 |
qed "part_ord_0"; |
|
167 |
||
168 |
goalw Order.thy [linear_def] "linear(0,r)"; |
|
169 |
by (fast_tac ZF_cs 1); |
|
170 |
qed "linear_0"; |
|
171 |
||
172 |
goalw Order.thy [tot_ord_def] "tot_ord(0,r)"; |
|
173 |
by (simp_tac (ZF_ss addsimps [part_ord_0, linear_0]) 1); |
|
174 |
qed "tot_ord_0"; |
|
175 |
||
176 |
goalw Order.thy [wf_on_def, wf_def] "wf[0](r)"; |
|
177 |
by (fast_tac eq_cs 1); |
|
178 |
qed "wf_on_0"; |
|
179 |
||
180 |
goalw Order.thy [well_ord_def] "well_ord(0,r)"; |
|
181 |
by (simp_tac (ZF_ss addsimps [tot_ord_0, wf_on_0]) 1); |
|
182 |
qed "well_ord_0"; |
|
836
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
183 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
184 |
|
789 | 185 |
(** Order-preserving (monotone) maps **) |
186 |
||
187 |
goalw Order.thy [mono_map_def] |
|
188 |
"!!f. f: mono_map(A,r,B,s) ==> f: A->B"; |
|
189 |
by (etac CollectD1 1); |
|
190 |
qed "mono_map_is_fun"; |
|
191 |
||
192 |
goalw Order.thy [mono_map_def, inj_def] |
|
193 |
"!!f. [| linear(A,r); wf[B](s); f: mono_map(A,r,B,s) |] ==> f: inj(A,B)"; |
|
194 |
by (step_tac ZF_cs 1); |
|
195 |
by (linear_case_tac 1); |
|
196 |
by (REPEAT |
|
197 |
(EVERY [eresolve_tac [wf_on_not_refl RS notE] 1, |
|
1461 | 198 |
etac ssubst 2, |
199 |
fast_tac ZF_cs 2, |
|
200 |
REPEAT (ares_tac [apply_type] 1)])); |
|
789 | 201 |
qed "mono_map_is_inj"; |
202 |
||
203 |
||
204 |
(** Order-isomorphisms -- or similarities, as Suppes calls them **) |
|
205 |
||
848 | 206 |
val prems = goalw Order.thy [ord_iso_def] |
1461 | 207 |
"[| f: bij(A, B); \ |
208 |
\ !!x y. [| x:A; y:A |] ==> <x, y> : r <-> <f`x, f`y> : s \ |
|
848 | 209 |
\ |] ==> f: ord_iso(A,r,B,s)"; |
210 |
by (fast_tac (ZF_cs addSIs prems) 1); |
|
211 |
qed "ord_isoI"; |
|
212 |
||
789 | 213 |
goalw Order.thy [ord_iso_def, mono_map_def] |
214 |
"!!f. f: ord_iso(A,r,B,s) ==> f: mono_map(A,r,B,s)"; |
|
215 |
by (fast_tac (ZF_cs addSDs [bij_is_fun]) 1); |
|
216 |
qed "ord_iso_is_mono_map"; |
|
435 | 217 |
|
218 |
goalw Order.thy [ord_iso_def] |
|
219 |
"!!f. f: ord_iso(A,r,B,s) ==> f: bij(A,B)"; |
|
220 |
by (etac CollectD1 1); |
|
760 | 221 |
qed "ord_iso_is_bij"; |
435 | 222 |
|
789 | 223 |
(*Needed? But ord_iso_converse is!*) |
435 | 224 |
goalw Order.thy [ord_iso_def] |
225 |
"!!f. [| f: ord_iso(A,r,B,s); <x,y>: r; x:A; y:A |] ==> \ |
|
226 |
\ <f`x, f`y> : s"; |
|
227 |
by (fast_tac ZF_cs 1); |
|
760 | 228 |
qed "ord_iso_apply"; |
435 | 229 |
|
230 |
goalw Order.thy [ord_iso_def] |
|
231 |
"!!f. [| f: ord_iso(A,r,B,s); <x,y>: s; x:B; y:B |] ==> \ |
|
232 |
\ <converse(f) ` x, converse(f) ` y> : r"; |
|
437 | 233 |
by (etac CollectE 1); |
234 |
by (etac (bspec RS bspec RS iffD2) 1); |
|
435 | 235 |
by (REPEAT (eresolve_tac [asm_rl, |
1461 | 236 |
bij_converse_bij RS bij_is_fun RS apply_type] 1)); |
435 | 237 |
by (asm_simp_tac (ZF_ss addsimps [right_inverse_bij]) 1); |
760 | 238 |
qed "ord_iso_converse"; |
435 | 239 |
|
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
240 |
|
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
241 |
(*Rewriting with bijections and converse (function inverse)*) |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
242 |
val bij_inverse_ss = |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
243 |
ZF_ss setsolver (type_auto_tac [ord_iso_is_bij, bij_is_fun, apply_type, |
1461 | 244 |
bij_converse_bij, comp_fun, comp_bij]) |
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
245 |
addsimps [right_inverse_bij, left_inverse_bij, comp_fun_apply]; |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
246 |
|
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
247 |
|
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
248 |
(** Symmetry and Transitivity Rules **) |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
249 |
|
848 | 250 |
(*Reflexivity of similarity*) |
251 |
goal Order.thy "id(A): ord_iso(A,r,A,r)"; |
|
252 |
by (resolve_tac [id_bij RS ord_isoI] 1); |
|
253 |
by (asm_simp_tac (ZF_ss addsimps [id_conv]) 1); |
|
254 |
qed "ord_iso_refl"; |
|
255 |
||
256 |
(*Symmetry of similarity*) |
|
769 | 257 |
goalw Order.thy [ord_iso_def] |
258 |
"!!f. f: ord_iso(A,r,B,s) ==> converse(f): ord_iso(B,s,A,r)"; |
|
990
9ec3c7bd774e
Tried the new addss in many proofs, and tidied others involving simplification.
lcp
parents:
848
diff
changeset
|
259 |
by (fast_tac (ZF_cs addss bij_inverse_ss) 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
769
diff
changeset
|
260 |
qed "ord_iso_sym"; |
769 | 261 |
|
848 | 262 |
(*Transitivity of similarity*) |
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
263 |
goalw Order.thy [mono_map_def] |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
264 |
"!!f. [| g: mono_map(A,r,B,s); f: mono_map(B,s,C,t) |] ==> \ |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
265 |
\ (f O g): mono_map(A,r,C,t)"; |
990
9ec3c7bd774e
Tried the new addss in many proofs, and tidied others involving simplification.
lcp
parents:
848
diff
changeset
|
266 |
by (fast_tac (ZF_cs addss bij_inverse_ss) 1); |
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
267 |
qed "mono_map_trans"; |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
268 |
|
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
269 |
(*Transitivity of similarity: the order-isomorphism relation*) |
789 | 270 |
goalw Order.thy [ord_iso_def] |
271 |
"!!f. [| g: ord_iso(A,r,B,s); f: ord_iso(B,s,C,t) |] ==> \ |
|
272 |
\ (f O g): ord_iso(A,r,C,t)"; |
|
990
9ec3c7bd774e
Tried the new addss in many proofs, and tidied others involving simplification.
lcp
parents:
848
diff
changeset
|
273 |
by (fast_tac (ZF_cs addss bij_inverse_ss) 1); |
789 | 274 |
qed "ord_iso_trans"; |
275 |
||
276 |
(** Two monotone maps can make an order-isomorphism **) |
|
277 |
||
278 |
goalw Order.thy [ord_iso_def, mono_map_def] |
|
1461 | 279 |
"!!f g. [| f: mono_map(A,r,B,s); g: mono_map(B,s,A,r); \ |
789 | 280 |
\ f O g = id(B); g O f = id(A) |] ==> f: ord_iso(A,r,B,s)"; |
281 |
by (safe_tac ZF_cs); |
|
282 |
by (REPEAT_FIRST (ares_tac [fg_imp_bijective])); |
|
283 |
by (fast_tac ZF_cs 1); |
|
284 |
by (subgoal_tac "<g`(f`x), g`(f`y)> : r" 1); |
|
285 |
by (fast_tac (ZF_cs addIs [apply_type] addSEs [bspec RS bspec RS mp]) 2); |
|
286 |
by (asm_full_simp_tac |
|
287 |
(ZF_ss addsimps [comp_eq_id_iff RS iffD1]) 1); |
|
288 |
qed "mono_ord_isoI"; |
|
289 |
||
290 |
goal Order.thy |
|
1461 | 291 |
"!!B. [| well_ord(A,r); well_ord(B,s); \ |
292 |
\ f: mono_map(A,r,B,s); converse(f): mono_map(B,s,A,r) \ |
|
789 | 293 |
\ |] ==> f: ord_iso(A,r,B,s)"; |
294 |
by (REPEAT (ares_tac [mono_ord_isoI] 1)); |
|
295 |
by (forward_tac [mono_map_is_fun RS fun_is_rel] 1); |
|
296 |
by (etac (converse_converse RS subst) 1 THEN rtac left_comp_inverse 1); |
|
297 |
by (DEPTH_SOLVE (ares_tac [mono_map_is_inj, left_comp_inverse] 1 |
|
298 |
ORELSE eresolve_tac [well_ord_is_linear, well_ord_is_wf] 1)); |
|
299 |
qed "well_ord_mono_ord_isoI"; |
|
300 |
||
301 |
||
302 |
(** Order-isomorphisms preserve the ordering's properties **) |
|
303 |
||
304 |
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def, ord_iso_def] |
|
305 |
"!!A B r. [| part_ord(B,s); f: ord_iso(A,r,B,s) |] ==> part_ord(A,r)"; |
|
306 |
by (asm_simp_tac ZF_ss 1); |
|
307 |
by (fast_tac (ZF_cs addIs [bij_is_fun RS apply_type]) 1); |
|
308 |
qed "part_ord_ord_iso"; |
|
435 | 309 |
|
789 | 310 |
goalw Order.thy [linear_def, ord_iso_def] |
311 |
"!!A B r. [| linear(B,s); f: ord_iso(A,r,B,s) |] ==> linear(A,r)"; |
|
312 |
by (asm_simp_tac ZF_ss 1); |
|
313 |
by (safe_tac ZF_cs); |
|
314 |
by (dres_inst_tac [("x1", "f`x"), ("x", "f`xa")] (bspec RS bspec) 1); |
|
315 |
by (safe_tac (ZF_cs addSEs [bij_is_fun RS apply_type])); |
|
316 |
by (dres_inst_tac [("t", "op `(converse(f))")] subst_context 1); |
|
317 |
by (asm_full_simp_tac (ZF_ss addsimps [left_inverse_bij]) 1); |
|
318 |
qed "linear_ord_iso"; |
|
319 |
||
320 |
goalw Order.thy [wf_on_def, wf_def, ord_iso_def] |
|
321 |
"!!A B r. [| wf[B](s); f: ord_iso(A,r,B,s) |] ==> wf[A](r)"; |
|
322 |
(*reversed &-congruence rule handles context of membership in A*) |
|
323 |
by (asm_full_simp_tac (ZF_ss addcongs [conj_cong2]) 1); |
|
324 |
by (safe_tac ZF_cs); |
|
325 |
by (dres_inst_tac [("x", "{f`z. z:Z Int A}")] spec 1); |
|
326 |
by (safe_tac eq_cs); |
|
327 |
by (dtac equalityD1 1); |
|
328 |
by (fast_tac (ZF_cs addSIs [bexI]) 1); |
|
329 |
by (fast_tac (ZF_cs addSIs [bexI] |
|
1461 | 330 |
addIs [bij_is_fun RS apply_type]) 1); |
789 | 331 |
qed "wf_on_ord_iso"; |
332 |
||
333 |
goalw Order.thy [well_ord_def, tot_ord_def] |
|
334 |
"!!A B r. [| well_ord(B,s); f: ord_iso(A,r,B,s) |] ==> well_ord(A,r)"; |
|
335 |
by (fast_tac |
|
336 |
(ZF_cs addSEs [part_ord_ord_iso, linear_ord_iso, wf_on_ord_iso]) 1); |
|
337 |
qed "well_ord_ord_iso"; |
|
338 |
||
339 |
||
340 |
(*** Main results of Kunen, Chapter 1 section 6 ***) |
|
435 | 341 |
|
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
342 |
(*Inductive argument for Kunen's Lemma 6.1, etc. |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
343 |
Simple proof from Halmos, page 72*) |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
344 |
goalw Order.thy [well_ord_def, ord_iso_def] |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
345 |
"!!r. [| well_ord(A,r); f: ord_iso(A,r, A',r); A'<= A; y: A |] \ |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
346 |
\ ==> ~ <f`y, y>: r"; |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
347 |
by (REPEAT (eresolve_tac [conjE, CollectE] 1)); |
435 | 348 |
by (wf_on_ind_tac "y" [] 1); |
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
349 |
by (dres_inst_tac [("a","y1")] (bij_is_fun RS apply_type) 1); |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
350 |
by (assume_tac 1); |
435 | 351 |
by (fast_tac ZF_cs 1); |
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
352 |
qed "well_ord_iso_subset_lemma"; |
435 | 353 |
|
354 |
(*Kunen's Lemma 6.1: there's no order-isomorphism to an initial segment |
|
355 |
of a well-ordering*) |
|
356 |
goal Order.thy |
|
789 | 357 |
"!!r. [| well_ord(A,r); f : ord_iso(A, r, pred(A,x,r), r); x:A |] ==> P"; |
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
358 |
by (metacut_tac well_ord_iso_subset_lemma 1); |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
359 |
by (REPEAT_FIRST (ares_tac [pred_subset])); |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
360 |
(*Now we know f`x < x *) |
435 | 361 |
by (EVERY1 [dtac (ord_iso_is_bij RS bij_is_fun RS apply_type), |
1461 | 362 |
assume_tac]); |
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
363 |
(*Now we also know f`x : pred(A,x,r); contradiction! *) |
435 | 364 |
by (asm_full_simp_tac (ZF_ss addsimps [well_ord_def, pred_def]) 1); |
789 | 365 |
qed "well_ord_iso_predE"; |
366 |
||
367 |
(*Simple consequence of Lemma 6.1*) |
|
368 |
goal Order.thy |
|
1461 | 369 |
"!!r. [| well_ord(A,r); f : ord_iso(pred(A,a,r), r, pred(A,c,r), r); \ |
789 | 370 |
\ a:A; c:A |] ==> a=c"; |
371 |
by (forward_tac [well_ord_is_trans_on] 1); |
|
372 |
by (forward_tac [well_ord_is_linear] 1); |
|
373 |
by (linear_case_tac 1); |
|
374 |
by (dtac ord_iso_sym 1); |
|
375 |
by (REPEAT (*because there are two symmetric cases*) |
|
376 |
(EVERY [eresolve_tac [pred_subset RSN (2, well_ord_subset) RS |
|
1461 | 377 |
well_ord_iso_predE] 1, |
378 |
fast_tac (ZF_cs addSIs [predI]) 2, |
|
379 |
asm_simp_tac (ZF_ss addsimps [trans_pred_pred_eq]) 1])); |
|
789 | 380 |
qed "well_ord_iso_pred_eq"; |
381 |
||
382 |
(*Does not assume r is a wellordering!*) |
|
383 |
goalw Order.thy [ord_iso_def, pred_def] |
|
1461 | 384 |
"!!r. [| f : ord_iso(A,r,B,s); a:A |] ==> \ |
789 | 385 |
\ f `` pred(A,a,r) = pred(B, f`a, s)"; |
386 |
by (etac CollectE 1); |
|
387 |
by (asm_simp_tac |
|
388 |
(ZF_ss addsimps [[bij_is_fun, Collect_subset] MRS image_fun]) 1); |
|
389 |
by (safe_tac (eq_cs addSEs [bij_is_fun RS apply_type])); |
|
1461 | 390 |
by (rtac RepFun_eqI 1); |
789 | 391 |
by (fast_tac (ZF_cs addSIs [right_inverse_bij RS sym]) 1); |
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
392 |
by (asm_simp_tac bij_inverse_ss 1); |
789 | 393 |
qed "ord_iso_image_pred"; |
435 | 394 |
|
789 | 395 |
(*But in use, A and B may themselves be initial segments. Then use |
396 |
trans_pred_pred_eq to simplify the pred(pred...) terms. See just below.*) |
|
397 |
goal Order.thy |
|
1461 | 398 |
"!!r. [| f : ord_iso(A,r,B,s); a:A |] ==> \ |
789 | 399 |
\ restrict(f, pred(A,a,r)) : ord_iso(pred(A,a,r), r, pred(B, f`a, s), s)"; |
400 |
by (asm_simp_tac (ZF_ss addsimps [ord_iso_image_pred RS sym]) 1); |
|
1461 | 401 |
by (rewtac ord_iso_def); |
848 | 402 |
by (etac CollectE 1); |
403 |
by (rtac CollectI 1); |
|
789 | 404 |
by (asm_full_simp_tac (ZF_ss addsimps [pred_def]) 2); |
405 |
by (eresolve_tac [[bij_is_inj, pred_subset] MRS restrict_bij] 1); |
|
406 |
qed "ord_iso_restrict_pred"; |
|
407 |
||
408 |
(*Tricky; a lot of forward proof!*) |
|
409 |
goal Order.thy |
|
1461 | 410 |
"!!r. [| well_ord(A,r); well_ord(B,s); <a,c>: r; \ |
411 |
\ f : ord_iso(pred(A,a,r), r, pred(B,b,s), s); \ |
|
412 |
\ g : ord_iso(pred(A,c,r), r, pred(B,d,s), s); \ |
|
789 | 413 |
\ a:A; c:A; b:B; d:B |] ==> <b,d>: s"; |
414 |
by (forward_tac [ord_iso_is_bij RS bij_is_fun RS apply_type] 1 THEN |
|
415 |
REPEAT1 (eresolve_tac [asm_rl, predI, predE] 1)); |
|
416 |
by (subgoal_tac "b = g`a" 1); |
|
417 |
by (asm_simp_tac ZF_ss 1); |
|
1461 | 418 |
by (rtac well_ord_iso_pred_eq 1); |
789 | 419 |
by (REPEAT_SOME assume_tac); |
420 |
by (forward_tac [ord_iso_restrict_pred] 1 THEN |
|
421 |
REPEAT1 (eresolve_tac [asm_rl, predI] 1)); |
|
422 |
by (asm_full_simp_tac |
|
423 |
(ZF_ss addsimps [well_ord_is_trans_on, trans_pred_pred_eq]) 1); |
|
424 |
by (eresolve_tac [ord_iso_sym RS ord_iso_trans] 1); |
|
425 |
by (assume_tac 1); |
|
426 |
qed "well_ord_iso_preserving"; |
|
435 | 427 |
|
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
428 |
(*See Halmos, page 72*) |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
429 |
goal Order.thy |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
430 |
"!!r. [| well_ord(A,r); \ |
435 | 431 |
\ f: ord_iso(A,r, B,s); g: ord_iso(A,r, B,s); y: A |] \ |
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
432 |
\ ==> ~ <g`y, f`y> : s"; |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
433 |
by (forward_tac [well_ord_iso_subset_lemma] 1); |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
434 |
by (res_inst_tac [("f","converse(f)"), ("g","g")] ord_iso_trans 1); |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
435 |
by (REPEAT_FIRST (ares_tac [subset_refl, ord_iso_sym])); |
435 | 436 |
by (safe_tac ZF_cs); |
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
437 |
by (forward_tac [ord_iso_converse] 1); |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
438 |
by (REPEAT (fast_tac (bij_apply_cs addSEs [ord_iso_is_bij]) 1)); |
435 | 439 |
by (asm_full_simp_tac bij_inverse_ss 1); |
760 | 440 |
qed "well_ord_iso_unique_lemma"; |
435 | 441 |
|
442 |
(*Kunen's Lemma 6.2: Order-isomorphisms between well-orderings are unique*) |
|
443 |
goal Order.thy |
|
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
444 |
"!!r. [| well_ord(A,r); \ |
435 | 445 |
\ f: ord_iso(A,r, B,s); g: ord_iso(A,r, B,s) |] ==> f = g"; |
437 | 446 |
by (rtac fun_extension 1); |
812
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
447 |
by (REPEAT (etac (ord_iso_is_bij RS bij_is_fun) 1)); |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
448 |
by (subgoals_tac ["f`x : B", "g`x : B", "linear(B,s)"] 1); |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
449 |
by (REPEAT (fast_tac (bij_apply_cs addSEs [ord_iso_is_bij]) 3)); |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
450 |
by (dtac well_ord_ord_iso 2 THEN etac ord_iso_sym 2); |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
451 |
by (asm_full_simp_tac (ZF_ss addsimps [tot_ord_def, well_ord_def]) 2); |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
452 |
by (linear_case_tac 1); |
bf4b7c37db2c
Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents:
794
diff
changeset
|
453 |
by (DEPTH_SOLVE (eresolve_tac [asm_rl, well_ord_iso_unique_lemma RS notE] 1)); |
760 | 454 |
qed "well_ord_iso_unique"; |
435 | 455 |
|
456 |
||
789 | 457 |
(** Towards Kunen's Theorem 6.3: linearity of the similarity relation **) |
458 |
||
459 |
goalw Order.thy [ord_iso_map_def] |
|
460 |
"ord_iso_map(A,r,B,s) <= A*B"; |
|
461 |
by (fast_tac ZF_cs 1); |
|
462 |
qed "ord_iso_map_subset"; |
|
463 |
||
464 |
goalw Order.thy [ord_iso_map_def] |
|
465 |
"domain(ord_iso_map(A,r,B,s)) <= A"; |
|
466 |
by (fast_tac ZF_cs 1); |
|
467 |
qed "domain_ord_iso_map"; |
|
468 |
||
469 |
goalw Order.thy [ord_iso_map_def] |
|
470 |
"range(ord_iso_map(A,r,B,s)) <= B"; |
|
471 |
by (fast_tac ZF_cs 1); |
|
472 |
qed "range_ord_iso_map"; |
|
473 |
||
474 |
goalw Order.thy [ord_iso_map_def] |
|
475 |
"converse(ord_iso_map(A,r,B,s)) = ord_iso_map(B,s,A,r)"; |
|
476 |
by (fast_tac (eq_cs addIs [ord_iso_sym]) 1); |
|
477 |
qed "converse_ord_iso_map"; |
|
478 |
||
479 |
goalw Order.thy [ord_iso_map_def, function_def] |
|
480 |
"!!B. well_ord(B,s) ==> function(ord_iso_map(A,r,B,s))"; |
|
435 | 481 |
by (safe_tac ZF_cs); |
1461 | 482 |
by (rtac well_ord_iso_pred_eq 1); |
789 | 483 |
by (REPEAT_SOME assume_tac); |
484 |
by (eresolve_tac [ord_iso_sym RS ord_iso_trans] 1); |
|
485 |
by (assume_tac 1); |
|
486 |
qed "function_ord_iso_map"; |
|
487 |
||
488 |
goal Order.thy |
|
1461 | 489 |
"!!B. well_ord(B,s) ==> ord_iso_map(A,r,B,s) \ |
789 | 490 |
\ : domain(ord_iso_map(A,r,B,s)) -> range(ord_iso_map(A,r,B,s))"; |
491 |
by (asm_simp_tac |
|
492 |
(ZF_ss addsimps [Pi_iff, function_ord_iso_map, |
|
1461 | 493 |
ord_iso_map_subset RS domain_times_range]) 1); |
789 | 494 |
qed "ord_iso_map_fun"; |
435 | 495 |
|
789 | 496 |
goalw Order.thy [mono_map_def] |
1461 | 497 |
"!!B. [| well_ord(A,r); well_ord(B,s) |] ==> ord_iso_map(A,r,B,s) \ |
498 |
\ : mono_map(domain(ord_iso_map(A,r,B,s)), r, \ |
|
499 |
\ range(ord_iso_map(A,r,B,s)), s)"; |
|
789 | 500 |
by (asm_simp_tac (ZF_ss addsimps [ord_iso_map_fun]) 1); |
435 | 501 |
by (safe_tac ZF_cs); |
789 | 502 |
by (subgoals_tac ["x:A", "xa:A", "y:B", "ya:B"] 1); |
503 |
by (REPEAT |
|
504 |
(fast_tac (ZF_cs addSEs [ord_iso_map_subset RS subsetD RS SigmaE]) 2)); |
|
505 |
by (asm_simp_tac (ZF_ss addsimps [ord_iso_map_fun RSN (2,apply_equality)]) 1); |
|
1461 | 506 |
by (rewtac ord_iso_map_def); |
789 | 507 |
by (safe_tac (ZF_cs addSEs [UN_E])); |
1461 | 508 |
by (rtac well_ord_iso_preserving 1 THEN REPEAT_FIRST assume_tac); |
789 | 509 |
qed "ord_iso_map_mono_map"; |
435 | 510 |
|
789 | 511 |
goalw Order.thy [mono_map_def] |
1461 | 512 |
"!!B. [| well_ord(A,r); well_ord(B,s) |] ==> ord_iso_map(A,r,B,s) \ |
513 |
\ : ord_iso(domain(ord_iso_map(A,r,B,s)), r, \ |
|
514 |
\ range(ord_iso_map(A,r,B,s)), s)"; |
|
515 |
by (rtac well_ord_mono_ord_isoI 1); |
|
789 | 516 |
by (resolve_tac [converse_ord_iso_map RS subst] 4); |
517 |
by (asm_simp_tac |
|
518 |
(ZF_ss addsimps [ord_iso_map_subset RS converse_converse, |
|
1461 | 519 |
domain_converse, range_converse]) 4); |
789 | 520 |
by (REPEAT (ares_tac [ord_iso_map_mono_map] 3)); |
521 |
by (ALLGOALS (etac well_ord_subset)); |
|
522 |
by (ALLGOALS (resolve_tac [domain_ord_iso_map, range_ord_iso_map])); |
|
523 |
qed "ord_iso_map_ord_iso"; |
|
467 | 524 |
|
789 | 525 |
(*One way of saying that domain(ord_iso_map(A,r,B,s)) is downwards-closed*) |
526 |
goalw Order.thy [ord_iso_map_def] |
|
1461 | 527 |
"!!B. [| well_ord(A,r); well_ord(B,s); \ |
528 |
\ a: A; a ~: domain(ord_iso_map(A,r,B,s)) \ |
|
1015
75110179587d
Changed proof of domain_ord_iso_map_subset for new hyp_subst_tac
lcp
parents:
990
diff
changeset
|
529 |
\ |] ==> domain(ord_iso_map(A,r,B,s)) <= pred(A, a, r)"; |
75110179587d
Changed proof of domain_ord_iso_map_subset for new hyp_subst_tac
lcp
parents:
990
diff
changeset
|
530 |
by (safe_tac (ZF_cs addSIs [predI])); |
75110179587d
Changed proof of domain_ord_iso_map_subset for new hyp_subst_tac
lcp
parents:
990
diff
changeset
|
531 |
(*Case analysis on xaa vs a in r *) |
789 | 532 |
by (forw_inst_tac [("A","A")] well_ord_is_linear 1); |
533 |
by (linear_case_tac 1); |
|
1015
75110179587d
Changed proof of domain_ord_iso_map_subset for new hyp_subst_tac
lcp
parents:
990
diff
changeset
|
534 |
(*Trivial case: a=xa*) |
75110179587d
Changed proof of domain_ord_iso_map_subset for new hyp_subst_tac
lcp
parents:
990
diff
changeset
|
535 |
by (fast_tac ZF_cs 2); |
75110179587d
Changed proof of domain_ord_iso_map_subset for new hyp_subst_tac
lcp
parents:
990
diff
changeset
|
536 |
(*Harder case: <a, xa>: r*) |
789 | 537 |
by (forward_tac [ord_iso_is_bij RS bij_is_fun RS apply_type] 1 THEN |
538 |
REPEAT1 (eresolve_tac [asm_rl, predI, predE] 1)); |
|
539 |
by (forward_tac [ord_iso_restrict_pred] 1 THEN |
|
540 |
REPEAT1 (eresolve_tac [asm_rl, predI] 1)); |
|
541 |
by (asm_full_simp_tac |
|
542 |
(ZF_ss addsimps [well_ord_is_trans_on, trans_pred_pred_eq]) 1); |
|
543 |
by (fast_tac ZF_cs 1); |
|
544 |
qed "domain_ord_iso_map_subset"; |
|
435 | 545 |
|
789 | 546 |
(*For the 4-way case analysis in the main result*) |
547 |
goal Order.thy |
|
548 |
"!!B. [| well_ord(A,r); well_ord(B,s) |] ==> \ |
|
1461 | 549 |
\ domain(ord_iso_map(A,r,B,s)) = A | \ |
789 | 550 |
\ (EX x:A. domain(ord_iso_map(A,r,B,s)) = pred(A,x,r))"; |
551 |
by (forward_tac [well_ord_is_wf] 1); |
|
552 |
by (rewrite_goals_tac [wf_on_def, wf_def]); |
|
553 |
by (dres_inst_tac [("x", "A-domain(ord_iso_map(A,r,B,s))")] spec 1); |
|
554 |
by (step_tac ZF_cs 1); |
|
555 |
(*The first case: the domain equals A*) |
|
556 |
by (rtac (domain_ord_iso_map RS equalityI) 1); |
|
557 |
by (etac (Diff_eq_0_iff RS iffD1) 1); |
|
558 |
(*The other case: the domain equals an initial segment*) |
|
559 |
by (swap_res_tac [bexI] 1); |
|
560 |
by (assume_tac 2); |
|
561 |
by (rtac equalityI 1); |
|
562 |
(*not ZF_cs below; that would use rules like domainE!*) |
|
563 |
by (fast_tac (pair_cs addSEs [predE]) 2); |
|
564 |
by (REPEAT (ares_tac [domain_ord_iso_map_subset] 1)); |
|
565 |
qed "domain_ord_iso_map_cases"; |
|
467 | 566 |
|
789 | 567 |
(*As above, by duality*) |
568 |
goal Order.thy |
|
1461 | 569 |
"!!B. [| well_ord(A,r); well_ord(B,s) |] ==> \ |
570 |
\ range(ord_iso_map(A,r,B,s)) = B | \ |
|
789 | 571 |
\ (EX y:B. range(ord_iso_map(A,r,B,s))= pred(B,y,s))"; |
572 |
by (resolve_tac [converse_ord_iso_map RS subst] 1); |
|
573 |
by (asm_simp_tac |
|
574 |
(ZF_ss addsimps [range_converse, domain_ord_iso_map_cases]) 1); |
|
575 |
qed "range_ord_iso_map_cases"; |
|
576 |
||
577 |
(*Kunen's Theorem 6.3: Fundamental Theorem for Well-Ordered Sets*) |
|
578 |
goal Order.thy |
|
1461 | 579 |
"!!B. [| well_ord(A,r); well_ord(B,s) |] ==> \ |
580 |
\ ord_iso_map(A,r,B,s) : ord_iso(A, r, B, s) | \ |
|
789 | 581 |
\ (EX x:A. ord_iso_map(A,r,B,s) : ord_iso(pred(A,x,r), r, B, s)) | \ |
582 |
\ (EX y:B. ord_iso_map(A,r,B,s) : ord_iso(A, r, pred(B,y,s), s))"; |
|
583 |
by (forw_inst_tac [("B","B")] domain_ord_iso_map_cases 1); |
|
584 |
by (forw_inst_tac [("B","B")] range_ord_iso_map_cases 2); |
|
585 |
by (REPEAT_FIRST (eresolve_tac [asm_rl, disjE, bexE])); |
|
1461 | 586 |
by (ALLGOALS (dtac ord_iso_map_ord_iso THEN' assume_tac THEN' |
587 |
asm_full_simp_tac (ZF_ss addsimps [bexI]))); |
|
789 | 588 |
by (resolve_tac [wf_on_not_refl RS notE] 1); |
1461 | 589 |
by (etac well_ord_is_wf 1); |
789 | 590 |
by (assume_tac 1); |
591 |
by (subgoal_tac "<x,y>: ord_iso_map(A,r,B,s)" 1); |
|
1461 | 592 |
by (dtac rangeI 1); |
789 | 593 |
by (asm_full_simp_tac (ZF_ss addsimps [pred_def]) 1); |
1461 | 594 |
by (rewtac ord_iso_map_def); |
467 | 595 |
by (fast_tac ZF_cs 1); |
789 | 596 |
qed "well_ord_trichotomy"; |
467 | 597 |
|
836
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
598 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
599 |
(*** Properties of converse(r), by Krzysztof Grabczewski ***) |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
600 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
601 |
goalw Order.thy [irrefl_def] |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
602 |
"!!A. irrefl(A,r) ==> irrefl(A,converse(r))"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
603 |
by (fast_tac (ZF_cs addSIs [converseI]) 1); |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
604 |
qed "irrefl_converse"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
605 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
606 |
goalw Order.thy [trans_on_def] |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
607 |
"!!A. trans[A](r) ==> trans[A](converse(r))"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
608 |
by (fast_tac (ZF_cs addSIs [converseI]) 1); |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
609 |
qed "trans_on_converse"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
610 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
611 |
goalw Order.thy [part_ord_def] |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
612 |
"!!A. part_ord(A,r) ==> part_ord(A,converse(r))"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
613 |
by (fast_tac (ZF_cs addSIs [irrefl_converse, trans_on_converse]) 1); |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
614 |
qed "part_ord_converse"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
615 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
616 |
goalw Order.thy [linear_def] |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
617 |
"!!A. linear(A,r) ==> linear(A,converse(r))"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
618 |
by (fast_tac (ZF_cs addSIs [converseI]) 1); |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
619 |
qed "linear_converse"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
620 |
|
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
621 |
goalw Order.thy [tot_ord_def] |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
622 |
"!!A. tot_ord(A,r) ==> tot_ord(A,converse(r))"; |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
623 |
by (fast_tac (ZF_cs addSIs [part_ord_converse, linear_converse]) 1); |
627f4842b020
Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents:
812
diff
changeset
|
624 |
qed "tot_ord_converse"; |