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