435
|
1 |
(* Title: ZF/Order.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1994 University of Cambridge
|
|
5 |
|
|
6 |
For Order.thy. Orders in Zermelo-Fraenkel Set Theory
|
|
7 |
*)
|
|
8 |
|
|
9 |
|
|
10 |
open Order;
|
|
11 |
|
|
12 |
val bij_apply_cs = ZF_cs addSEs [bij_converse_bij]
|
|
13 |
addIs [bij_is_fun, apply_type];
|
|
14 |
|
|
15 |
val bij_inverse_ss =
|
|
16 |
ZF_ss addsimps [bij_is_fun RS apply_type,
|
|
17 |
bij_converse_bij RS bij_is_fun RS apply_type,
|
|
18 |
left_inverse_bij, right_inverse_bij];
|
|
19 |
|
|
20 |
(** Basic properties of the definitions **)
|
|
21 |
|
|
22 |
(*needed????*)
|
|
23 |
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def, asym_def]
|
|
24 |
"!!r. part_ord(A,r) ==> asym(r Int A*A)";
|
|
25 |
by (fast_tac ZF_cs 1);
|
|
26 |
val part_ord_Imp_asym = result();
|
|
27 |
|
484
|
28 |
(** Subset properties for the various forms of relation **)
|
|
29 |
|
|
30 |
|
|
31 |
(*Note: a relation s such that s<=r need not be a partial ordering*)
|
|
32 |
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def]
|
|
33 |
"!!A B r. [| part_ord(A,r); B<=A |] ==> part_ord(B,r)";
|
|
34 |
by (fast_tac ZF_cs 1);
|
|
35 |
val part_ord_subset = result();
|
|
36 |
|
|
37 |
goalw Order.thy [linear_def]
|
|
38 |
"!!A B r. [| linear(A,r); B<=A |] ==> linear(B,r)";
|
|
39 |
by (fast_tac ZF_cs 1);
|
|
40 |
val linear_subset = result();
|
|
41 |
|
|
42 |
goalw Order.thy [tot_ord_def]
|
|
43 |
"!!A B r. [| tot_ord(A,r); B<=A |] ==> tot_ord(B,r)";
|
|
44 |
by (fast_tac (ZF_cs addSEs [part_ord_subset, linear_subset]) 1);
|
|
45 |
val tot_ord_subset = result();
|
|
46 |
|
|
47 |
goalw Order.thy [well_ord_def]
|
|
48 |
"!!A B r. [| well_ord(A,r); B<=A |] ==> well_ord(B,r)";
|
|
49 |
by (fast_tac (ZF_cs addSEs [tot_ord_subset, wf_on_subset_A]) 1);
|
|
50 |
val well_ord_subset = result();
|
|
51 |
|
|
52 |
|
435
|
53 |
(** Order-isomorphisms **)
|
|
54 |
|
|
55 |
goalw Order.thy [ord_iso_def]
|
|
56 |
"!!f. f: ord_iso(A,r,B,s) ==> f: bij(A,B)";
|
|
57 |
by (etac CollectD1 1);
|
|
58 |
val ord_iso_is_bij = result();
|
|
59 |
|
|
60 |
goalw Order.thy [ord_iso_def]
|
|
61 |
"!!f. [| f: ord_iso(A,r,B,s); <x,y>: r; x:A; y:A |] ==> \
|
|
62 |
\ <f`x, f`y> : s";
|
|
63 |
by (fast_tac ZF_cs 1);
|
|
64 |
val ord_iso_apply = result();
|
|
65 |
|
|
66 |
goalw Order.thy [ord_iso_def]
|
|
67 |
"!!f. [| f: ord_iso(A,r,B,s); <x,y>: s; x:B; y:B |] ==> \
|
|
68 |
\ <converse(f) ` x, converse(f) ` y> : r";
|
437
|
69 |
by (etac CollectE 1);
|
|
70 |
by (etac (bspec RS bspec RS iffD2) 1);
|
435
|
71 |
by (REPEAT (eresolve_tac [asm_rl,
|
|
72 |
bij_converse_bij RS bij_is_fun RS apply_type] 1));
|
|
73 |
by (asm_simp_tac (ZF_ss addsimps [right_inverse_bij]) 1);
|
|
74 |
val ord_iso_converse = result();
|
|
75 |
|
|
76 |
val major::premx::premy::prems = goalw Order.thy [linear_def]
|
|
77 |
"[| linear(A,r); x:A; y:A; \
|
|
78 |
\ <x,y>:r ==> P; x=y ==> P; <y,x>:r ==> P |] ==> P";
|
|
79 |
by (cut_facts_tac [major,premx,premy] 1);
|
|
80 |
by (REPEAT_FIRST (eresolve_tac [ballE,disjE]));
|
|
81 |
by (EVERY1 (map etac prems));
|
|
82 |
by (ALLGOALS contr_tac);
|
|
83 |
val linearE = result();
|
|
84 |
|
|
85 |
(*Does the case analysis, deleting linear(A,r) and proving trivial subgoals*)
|
|
86 |
val linear_case_tac =
|
|
87 |
SELECT_GOAL (EVERY [etac linearE 1, assume_tac 1, assume_tac 1,
|
|
88 |
REPEAT_SOME assume_tac]);
|
|
89 |
|
|
90 |
(*Inductive argument for proving Kunen's Lemma 6.1*)
|
|
91 |
goalw Order.thy [well_ord_def, tot_ord_def, part_ord_def, pred_def]
|
|
92 |
"!!r. [| well_ord(A,r); x: A; f: ord_iso(A, r, pred(A,x,r), r); y: A |] \
|
|
93 |
\ ==> f`y = y";
|
|
94 |
by (safe_tac ZF_cs);
|
|
95 |
by (wf_on_ind_tac "y" [] 1);
|
|
96 |
by (forward_tac [ord_iso_is_bij] 1);
|
|
97 |
by (subgoal_tac "f`y1 : {y: A . <y, x> : r}" 1);
|
|
98 |
by (fast_tac (ZF_cs addSEs [bij_is_fun RS apply_type]) 2);
|
|
99 |
(*Now we know f`y1 : A and <f`y1, x> : r *)
|
|
100 |
by (etac CollectE 1);
|
|
101 |
by (linear_case_tac 1);
|
|
102 |
(*Case <f`y1, y1> : r *) (*use induction hyp*)
|
|
103 |
by (dtac bspec 1 THEN mp_tac 2 THEN assume_tac 1);
|
|
104 |
by (dres_inst_tac [("t", "op `(converse(f))")] subst_context 1);
|
|
105 |
by (asm_full_simp_tac (ZF_ss addsimps [left_inverse_bij]) 1);
|
|
106 |
(*The case <y1, f`y1> : r *)
|
|
107 |
by (subgoal_tac "<y1,x> : r" 1);
|
|
108 |
by (fast_tac (ZF_cs addSEs [trans_onD]) 2);
|
|
109 |
by (dtac ord_iso_converse 1 THEN assume_tac 1 THEN fast_tac ZF_cs 2);
|
|
110 |
by (fast_tac ZF_cs 1);
|
|
111 |
by (asm_full_simp_tac (ZF_ss addsimps [left_inverse_bij]) 1);
|
|
112 |
(*now use induction hyp*)
|
|
113 |
by (dtac bspec 1 THEN mp_tac 2 THEN fast_tac bij_apply_cs 1);
|
|
114 |
by (dres_inst_tac [("t", "op `(f)")] subst_context 1);
|
|
115 |
by (asm_full_simp_tac (ZF_ss addsimps [right_inverse_bij]) 1);
|
|
116 |
val not_well_ord_iso_pred_lemma = result();
|
|
117 |
|
|
118 |
|
|
119 |
(*Kunen's Lemma 6.1: there's no order-isomorphism to an initial segment
|
|
120 |
of a well-ordering*)
|
|
121 |
goal Order.thy
|
|
122 |
"!!r. [| well_ord(A,r); x:A |] ==> \
|
|
123 |
\ ALL f. f ~: ord_iso(A, r, pred(A,x,r), r)";
|
|
124 |
by (safe_tac ZF_cs);
|
|
125 |
by (metacut_tac not_well_ord_iso_pred_lemma 1);
|
|
126 |
by (REPEAT_SOME assume_tac);
|
|
127 |
(*Now we know f`x = x*)
|
|
128 |
by (EVERY1 [dtac (ord_iso_is_bij RS bij_is_fun RS apply_type),
|
|
129 |
assume_tac]);
|
|
130 |
(*Now we know f`x : pred(A,x,r) *)
|
|
131 |
by (asm_full_simp_tac (ZF_ss addsimps [well_ord_def, pred_def]) 1);
|
|
132 |
by (fast_tac (ZF_cs addSEs [wf_on_not_refl RS notE]) 1);
|
|
133 |
val not_well_ord_iso_pred = result();
|
|
134 |
|
|
135 |
|
|
136 |
(*Inductive argument for proving Kunen's Lemma 6.2*)
|
|
137 |
goalw Order.thy [well_ord_def, tot_ord_def, part_ord_def]
|
|
138 |
"!!r. [| well_ord(A,r); well_ord(B,s); \
|
|
139 |
\ f: ord_iso(A,r, B,s); g: ord_iso(A,r, B,s); y: A |] \
|
|
140 |
\ ==> f`y = g`y";
|
|
141 |
by (safe_tac ZF_cs);
|
|
142 |
by (wf_on_ind_tac "y" [] 1);
|
|
143 |
by (subgoals_tac ["f: bij(A,B)", "g: bij(A,B)", "f`y1 : B", "g`y1 : B"] 1);
|
|
144 |
by (REPEAT (fast_tac (bij_apply_cs addSEs [ord_iso_is_bij]) 2));
|
|
145 |
by (linear_case_tac 1);
|
|
146 |
(*The case <f`y1, g`y1> : s *)
|
|
147 |
by (forw_inst_tac [("f","g")] ord_iso_converse 1 THEN REPEAT (assume_tac 1));
|
|
148 |
by (asm_full_simp_tac (ZF_ss addsimps [left_inverse_bij]) 1);
|
|
149 |
by (dtac bspec 1 THEN mp_tac 2 THEN fast_tac bij_apply_cs 1);
|
|
150 |
by (dres_inst_tac [("t", "op `(converse(f))")] subst_context 1);
|
|
151 |
by (dres_inst_tac [("t", "op `(g)")] subst_context 1);
|
|
152 |
by (asm_full_simp_tac bij_inverse_ss 1);
|
|
153 |
(*The case <g`y1, f`y1> : s *)
|
|
154 |
by (forw_inst_tac [("f","f")] ord_iso_converse 1 THEN REPEAT (assume_tac 1));
|
|
155 |
by (asm_full_simp_tac (ZF_ss addsimps [left_inverse_bij]) 1);
|
|
156 |
by (dtac bspec 1 THEN mp_tac 2 THEN fast_tac bij_apply_cs 1);
|
|
157 |
by (dres_inst_tac [("t", "op `(converse(g))")] subst_context 1);
|
|
158 |
by (dres_inst_tac [("t", "op `(f)")] subst_context 1);
|
|
159 |
by (asm_full_simp_tac bij_inverse_ss 1);
|
|
160 |
val well_ord_iso_unique_lemma = result();
|
|
161 |
|
|
162 |
|
|
163 |
(*Kunen's Lemma 6.2: Order-isomorphisms between well-orderings are unique*)
|
|
164 |
goal Order.thy
|
|
165 |
"!!r. [| well_ord(A,r); well_ord(B,s); \
|
|
166 |
\ f: ord_iso(A,r, B,s); g: ord_iso(A,r, B,s) |] ==> f = g";
|
437
|
167 |
by (rtac fun_extension 1);
|
|
168 |
by (etac (ord_iso_is_bij RS bij_is_fun) 1);
|
|
169 |
by (etac (ord_iso_is_bij RS bij_is_fun) 1);
|
|
170 |
by (rtac well_ord_iso_unique_lemma 1);
|
435
|
171 |
by (REPEAT_SOME assume_tac);
|
|
172 |
val well_ord_iso_unique = result();
|
|
173 |
|
|
174 |
|
|
175 |
goalw Order.thy [irrefl_def, part_ord_def, tot_ord_def,
|
|
176 |
trans_on_def, well_ord_def]
|
|
177 |
"!!r. [| wf[A](r); linear(A,r) |] ==> well_ord(A,r)";
|
|
178 |
by (asm_simp_tac (ZF_ss addsimps [wf_on_not_refl]) 1);
|
|
179 |
by (safe_tac ZF_cs);
|
|
180 |
by (linear_case_tac 1);
|
|
181 |
(*case x=xb*)
|
437
|
182 |
by (fast_tac (ZF_cs addSEs [wf_on_asym]) 1);
|
435
|
183 |
(*case x<xb*)
|
|
184 |
by (fast_tac (ZF_cs addSEs [wf_on_chain3]) 1);
|
|
185 |
val well_ordI = result();
|
|
186 |
|
|
187 |
goalw Order.thy [well_ord_def]
|
|
188 |
"!!r. well_ord(A,r) ==> wf[A](r)";
|
|
189 |
by (safe_tac ZF_cs);
|
|
190 |
val well_ord_is_wf = result();
|
|
191 |
|
467
|
192 |
goalw Order.thy [well_ord_def, tot_ord_def, part_ord_def]
|
|
193 |
"!!r. well_ord(A,r) ==> trans[A](r)";
|
|
194 |
by (safe_tac ZF_cs);
|
|
195 |
val well_ord_is_trans_on = result();
|
|
196 |
|
435
|
197 |
(*** Derived rules for pred(A,x,r) ***)
|
|
198 |
|
|
199 |
val [major,minor] = goalw Order.thy [pred_def]
|
|
200 |
"[| y: pred(A,x,r); [| y:A; <y,x>:r |] ==> P |] ==> P";
|
437
|
201 |
by (rtac (major RS CollectE) 1);
|
435
|
202 |
by (REPEAT (ares_tac [minor] 1));
|
|
203 |
val predE = result();
|
|
204 |
|
|
205 |
goalw Order.thy [pred_def] "pred(A,x,r) <= r -`` {x}";
|
|
206 |
by (fast_tac ZF_cs 1);
|
|
207 |
val pred_subset_under = result();
|
|
208 |
|
|
209 |
goalw Order.thy [pred_def] "pred(A,x,r) <= A";
|
|
210 |
by (fast_tac ZF_cs 1);
|
|
211 |
val pred_subset = result();
|
467
|
212 |
|
|
213 |
goalw Order.thy [pred_def] "y : pred(A,x,r) <-> <y,x>:r & y:A";
|
|
214 |
by (fast_tac ZF_cs 1);
|
|
215 |
val pred_iff = result();
|
|
216 |
|
|
217 |
goalw Order.thy [pred_def]
|
|
218 |
"pred(pred(A,x,r), y, r) = pred(A,x,r) Int pred(A,y,r)";
|
|
219 |
by (fast_tac eq_cs 1);
|
|
220 |
val pred_pred_eq = result();
|