author | wenzelm |
Wed, 11 Oct 2000 00:03:22 +0200 | |
changeset 10185 | c452fea3ce74 |
parent 9907 | 473a6604da94 |
child 12667 | 7e6eaaa125f2 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/OrderType.ML |
435 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
435 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
849 | 6 |
Order types and ordinal arithmetic in Zermelo-Fraenkel Set Theory |
7 |
||
8 |
Ordinal arithmetic is traditionally defined in terms of order types, as here. |
|
9 |
But a definition by transfinite recursion would be much simpler! |
|
435 | 10 |
*) |
11 |
||
12 |
||
849 | 13 |
(**** Proofs needing the combination of Ordinal.thy and Order.thy ****) |
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
14 |
|
9907 | 15 |
val [prem] = goal (the_context ()) "j le i ==> well_ord(j, Memrel(i))"; |
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
16 |
by (rtac well_ordI 1); |
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
17 |
by (rtac (wf_Memrel RS wf_imp_wf_on) 1); |
849 | 18 |
by (resolve_tac [prem RS ltE] 1); |
9842 | 19 |
by (asm_simp_tac (simpset() addsimps [linear_def, |
5268 | 20 |
[ltI, prem] MRS lt_trans2 RS ltD]) 1); |
849 | 21 |
by (REPEAT (resolve_tac [ballI, Ord_linear] 1)); |
22 |
by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 1)); |
|
23 |
qed "le_well_ord_Memrel"; |
|
24 |
||
25 |
(*"Ord(i) ==> well_ord(i, Memrel(i))"*) |
|
26 |
bind_thm ("well_ord_Memrel", le_refl RS le_well_ord_Memrel); |
|
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
27 |
|
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
28 |
(*Kunen's Theorem 7.3 (i), page 16; see also Ordinal/Ord_in_Ord |
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
29 |
The smaller ordinal is an initial segment of the larger *) |
5067 | 30 |
Goalw [pred_def, lt_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
31 |
"j<i ==> pred(i, j, Memrel(i)) = j"; |
9842 | 32 |
by (Asm_simp_tac 1); |
4091 | 33 |
by (blast_tac (claset() addIs [Ord_trans]) 1); |
849 | 34 |
qed "lt_pred_Memrel"; |
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
35 |
|
5067 | 36 |
Goalw [pred_def,Memrel_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
37 |
"x:A ==> pred(A, x, Memrel(A)) = A Int x"; |
2925 | 38 |
by (Blast_tac 1); |
831 | 39 |
qed "pred_Memrel"; |
40 |
||
5268 | 41 |
Goal "[| j<i; f: ord_iso(i,Memrel(i),j,Memrel(j)) |] ==> R"; |
7499 | 42 |
by (ftac lt_pred_Memrel 1); |
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
43 |
by (etac ltE 1); |
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
44 |
by (rtac (well_ord_Memrel RS well_ord_iso_predE) 1 THEN |
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
45 |
assume_tac 3 THEN assume_tac 1); |
9842 | 46 |
by (rewtac ord_iso_def); |
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
47 |
(*Combining the two simplifications causes looping*) |
9842 | 48 |
by (Asm_simp_tac 1); |
49 |
by (blast_tac (claset() addIs [bij_is_fun RS apply_type, Ord_trans]) 1); |
|
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
50 |
qed "Ord_iso_implies_eq_lemma"; |
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
51 |
|
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
52 |
(*Kunen's Theorem 7.3 (ii), page 16. Isomorphic ordinals are equal*) |
9842 | 53 |
Goal "[| Ord(i); Ord(j); f: ord_iso(i,Memrel(i),j,Memrel(j)) |] \ |
54 |
\ ==> i=j"; |
|
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
55 |
by (res_inst_tac [("i","i"),("j","j")] Ord_linear_lt 1); |
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
56 |
by (REPEAT (eresolve_tac [asm_rl, ord_iso_sym, Ord_iso_implies_eq_lemma] 1)); |
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
57 |
qed "Ord_iso_implies_eq"; |
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
58 |
|
849 | 59 |
|
60 |
(**** Ordermap and ordertype ****) |
|
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
61 |
|
5067 | 62 |
Goalw [ordermap_def,ordertype_def] |
437 | 63 |
"ordermap(A,r) : A -> ordertype(A,r)"; |
64 |
by (rtac lam_type 1); |
|
65 |
by (rtac (lamI RS imageI) 1); |
|
66 |
by (REPEAT (assume_tac 1)); |
|
760 | 67 |
qed "ordermap_type"; |
437 | 68 |
|
849 | 69 |
(*** Unfolding of ordermap ***) |
435 | 70 |
|
437 | 71 |
(*Useful for cardinality reasoning; see CardinalArith.ML*) |
5067 | 72 |
Goalw [ordermap_def, pred_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
73 |
"[| wf[A](r); x:A |] ==> \ |
437 | 74 |
\ ordermap(A,r) ` x = ordermap(A,r) `` pred(A,x,r)"; |
2469 | 75 |
by (Asm_simp_tac 1); |
437 | 76 |
by (etac (wfrec_on RS trans) 1); |
77 |
by (assume_tac 1); |
|
4091 | 78 |
by (asm_simp_tac (simpset() addsimps [subset_iff, image_lam, |
437 | 79 |
vimage_singleton_iff]) 1); |
760 | 80 |
qed "ordermap_eq_image"; |
437 | 81 |
|
467 | 82 |
(*Useful for rewriting PROVIDED pred is not unfolded until later!*) |
5268 | 83 |
Goal "[| wf[A](r); x:A |] ==> \ |
435 | 84 |
\ ordermap(A,r) ` x = {ordermap(A,r)`y . y : pred(A,x,r)}"; |
4091 | 85 |
by (asm_simp_tac (simpset() addsimps [ordermap_eq_image, pred_subset, |
1461 | 86 |
ordermap_type RS image_fun]) 1); |
760 | 87 |
qed "ordermap_pred_unfold"; |
435 | 88 |
|
89 |
(*pred-unfolded version. NOT suitable for rewriting -- loops!*) |
|
9907 | 90 |
bind_thm ("ordermap_unfold", rewrite_rule [pred_def] ordermap_pred_unfold); |
435 | 91 |
|
849 | 92 |
(*** Showing that ordermap, ordertype yield ordinals ***) |
435 | 93 |
|
94 |
fun ordermap_elim_tac i = |
|
95 |
EVERY [etac (ordermap_unfold RS equalityD1 RS subsetD RS RepFunE) i, |
|
1461 | 96 |
assume_tac (i+1), |
97 |
assume_tac i]; |
|
435 | 98 |
|
5067 | 99 |
Goalw [well_ord_def, tot_ord_def, part_ord_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
100 |
"[| well_ord(A,r); x:A |] ==> Ord(ordermap(A,r) ` x)"; |
4152 | 101 |
by Safe_tac; |
435 | 102 |
by (wf_on_ind_tac "x" [] 1); |
4091 | 103 |
by (asm_simp_tac (simpset() addsimps [ordermap_pred_unfold]) 1); |
435 | 104 |
by (rtac (Ord_is_Transset RSN (2,OrdI)) 1); |
437 | 105 |
by (rewrite_goals_tac [pred_def,Transset_def]); |
2925 | 106 |
by (Blast_tac 2); |
4152 | 107 |
by Safe_tac; |
435 | 108 |
by (ordermap_elim_tac 1); |
4091 | 109 |
by (fast_tac (claset() addSEs [trans_onD]) 1); |
760 | 110 |
qed "Ord_ordermap"; |
435 | 111 |
|
5067 | 112 |
Goalw [ordertype_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
113 |
"well_ord(A,r) ==> Ord(ordertype(A,r))"; |
2033 | 114 |
by (stac ([ordermap_type, subset_refl] MRS image_fun) 1); |
435 | 115 |
by (rtac (Ord_is_Transset RSN (2,OrdI)) 1); |
4091 | 116 |
by (blast_tac (claset() addIs [Ord_ordermap]) 2); |
437 | 117 |
by (rewrite_goals_tac [Transset_def,well_ord_def]); |
4152 | 118 |
by Safe_tac; |
435 | 119 |
by (ordermap_elim_tac 1); |
2925 | 120 |
by (Blast_tac 1); |
760 | 121 |
qed "Ord_ordertype"; |
435 | 122 |
|
849 | 123 |
(*** ordermap preserves the orderings in both directions ***) |
435 | 124 |
|
5268 | 125 |
Goal "[| <w,x>: r; wf[A](r); w: A; x: A |] ==> \ |
435 | 126 |
\ ordermap(A,r)`w : ordermap(A,r)`x"; |
127 |
by (eres_inst_tac [("x1", "x")] (ordermap_unfold RS ssubst) 1); |
|
437 | 128 |
by (assume_tac 1); |
2925 | 129 |
by (Blast_tac 1); |
760 | 130 |
qed "ordermap_mono"; |
435 | 131 |
|
132 |
(*linearity of r is crucial here*) |
|
5067 | 133 |
Goalw [well_ord_def, tot_ord_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
134 |
"[| ordermap(A,r)`w : ordermap(A,r)`x; well_ord(A,r); \ |
435 | 135 |
\ w: A; x: A |] ==> <w,x>: r"; |
4152 | 136 |
by Safe_tac; |
435 | 137 |
by (linear_case_tac 1); |
4091 | 138 |
by (blast_tac (claset() addSEs [mem_not_refl RS notE]) 1); |
467 | 139 |
by (dtac ordermap_mono 1); |
435 | 140 |
by (REPEAT_SOME assume_tac); |
437 | 141 |
by (etac mem_asym 1); |
142 |
by (assume_tac 1); |
|
760 | 143 |
qed "converse_ordermap_mono"; |
435 | 144 |
|
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
788
diff
changeset
|
145 |
bind_thm ("ordermap_surj", |
1461 | 146 |
rewrite_rule [symmetric ordertype_def] |
147 |
(ordermap_type RS surj_image)); |
|
435 | 148 |
|
5067 | 149 |
Goalw [well_ord_def, tot_ord_def, bij_def, inj_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
150 |
"well_ord(A,r) ==> ordermap(A,r) : bij(A, ordertype(A,r))"; |
4091 | 151 |
by (fast_tac (claset() addSIs [ordermap_type, ordermap_surj] |
3016 | 152 |
addEs [linearE] |
153 |
addDs [ordermap_mono] |
|
4091 | 154 |
addss (simpset() addsimps [mem_not_refl])) 1); |
760 | 155 |
qed "ordermap_bij"; |
435 | 156 |
|
849 | 157 |
(*** Isomorphisms involving ordertype ***) |
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
158 |
|
5067 | 159 |
Goalw [ord_iso_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
160 |
"well_ord(A,r) ==> \ |
435 | 161 |
\ ordermap(A,r) : ord_iso(A,r, ordertype(A,r), Memrel(ordertype(A,r)))"; |
4091 | 162 |
by (safe_tac (claset() addSEs [well_ord_is_wf] |
2925 | 163 |
addSIs [ordermap_type RS apply_type, |
164 |
ordermap_mono, ordermap_bij])); |
|
4091 | 165 |
by (blast_tac (claset() addSDs [converse_ordermap_mono]) 1); |
760 | 166 |
qed "ordertype_ord_iso"; |
435 | 167 |
|
5268 | 168 |
Goal "[| f: ord_iso(A,r,B,s); well_ord(B,s) |] ==> \ |
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
169 |
\ ordertype(A,r) = ordertype(B,s)"; |
7499 | 170 |
by (ftac well_ord_ord_iso 1 THEN assume_tac 1); |
1461 | 171 |
by (rtac Ord_iso_implies_eq 1 |
172 |
THEN REPEAT (etac Ord_ordertype 1)); |
|
4091 | 173 |
by (deepen_tac (claset() addIs [ord_iso_trans, ord_iso_sym] |
831 | 174 |
addSEs [ordertype_ord_iso]) 0 1); |
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
175 |
qed "ordertype_eq"; |
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
176 |
|
5268 | 177 |
Goal "[| ordertype(A,r) = ordertype(B,s); \ |
849 | 178 |
\ well_ord(A,r); well_ord(B,s) \ |
179 |
\ |] ==> EX f. f: ord_iso(A,r,B,s)"; |
|
1461 | 180 |
by (rtac exI 1); |
849 | 181 |
by (resolve_tac [ordertype_ord_iso RS ord_iso_trans] 1); |
182 |
by (assume_tac 1); |
|
1461 | 183 |
by (etac ssubst 1); |
849 | 184 |
by (eresolve_tac [ordertype_ord_iso RS ord_iso_sym] 1); |
185 |
qed "ordertype_eq_imp_ord_iso"; |
|
435 | 186 |
|
849 | 187 |
(*** Basic equalities for ordertype ***) |
467 | 188 |
|
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
189 |
(*Ordertype of Memrel*) |
5137 | 190 |
Goal "j le i ==> ordertype(j,Memrel(i)) = j"; |
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
191 |
by (resolve_tac [Ord_iso_implies_eq RS sym] 1); |
1461 | 192 |
by (etac ltE 1); |
849 | 193 |
by (REPEAT (ares_tac [le_well_ord_Memrel, Ord_ordertype] 1)); |
1461 | 194 |
by (rtac ord_iso_trans 1); |
849 | 195 |
by (eresolve_tac [le_well_ord_Memrel RS ordertype_ord_iso] 2); |
196 |
by (resolve_tac [id_bij RS ord_isoI] 1); |
|
9842 | 197 |
by (Asm_simp_tac 1); |
4091 | 198 |
by (fast_tac (claset() addEs [ltE, Ord_in_Ord, Ord_trans]) 1); |
849 | 199 |
qed "le_ordertype_Memrel"; |
200 |
||
201 |
(*"Ord(i) ==> ordertype(i, Memrel(i)) = i"*) |
|
202 |
bind_thm ("ordertype_Memrel", le_refl RS le_ordertype_Memrel); |
|
467 | 203 |
|
5067 | 204 |
Goal "ordertype(0,r) = 0"; |
849 | 205 |
by (resolve_tac [id_bij RS ord_isoI RS ordertype_eq RS trans] 1); |
206 |
by (etac emptyE 1); |
|
1461 | 207 |
by (rtac well_ord_0 1); |
849 | 208 |
by (resolve_tac [Ord_0 RS ordertype_Memrel] 1); |
209 |
qed "ordertype_0"; |
|
210 |
||
2469 | 211 |
Addsimps [ordertype_0]; |
212 |
||
849 | 213 |
(*Ordertype of rvimage: [| f: bij(A,B); well_ord(B,s) |] ==> |
214 |
ordertype(A, rvimage(A,f,s)) = ordertype(B,s) *) |
|
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
215 |
bind_thm ("bij_ordertype_vimage", ord_iso_rvimage RS ordertype_eq); |
467 | 216 |
|
849 | 217 |
(*** A fundamental unfolding law for ordertype. ***) |
218 |
||
814
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
lcp
parents:
807
diff
changeset
|
219 |
(*Ordermap returns the same result if applied to an initial segment*) |
5268 | 220 |
Goal "[| well_ord(A,r); y:A; z: pred(A,y,r) |] ==> \ |
1461 | 221 |
\ ordermap(pred(A,y,r), r) ` z = ordermap(A, r) ` z"; |
467 | 222 |
by (forward_tac [[well_ord_is_wf, pred_subset] MRS wf_on_subset_A] 1); |
223 |
by (wf_on_ind_tac "z" [] 1); |
|
4091 | 224 |
by (safe_tac (claset() addSEs [predE])); |
467 | 225 |
by (asm_simp_tac |
4091 | 226 |
(simpset() addsimps [ordermap_pred_unfold, well_ord_is_wf, pred_iff]) 1); |
467 | 227 |
(*combining these two simplifications LOOPS! *) |
4091 | 228 |
by (asm_simp_tac (simpset() addsimps [pred_pred_eq]) 1); |
229 |
by (asm_full_simp_tac (simpset() addsimps [pred_def]) 1); |
|
807 | 230 |
by (rtac (refl RSN (2,RepFun_cong)) 1); |
231 |
by (dtac well_ord_is_trans_on 1); |
|
4091 | 232 |
by (fast_tac (claset() addSEs [trans_onD]) 1); |
760 | 233 |
qed "ordermap_pred_eq_ordermap"; |
467 | 234 |
|
5067 | 235 |
Goalw [ordertype_def] |
849 | 236 |
"ordertype(A,r) = {ordermap(A,r)`y . y : A}"; |
237 |
by (rtac ([ordermap_type, subset_refl] MRS image_fun) 1); |
|
238 |
qed "ordertype_unfold"; |
|
239 |
||
240 |
(** Theorems by Krzysztof Grabczewski; proofs simplified by lcp **) |
|
241 |
||
5268 | 242 |
Goal "[| well_ord(A,r); x:A |] ==> \ |
849 | 243 |
\ ordertype(pred(A,x,r),r) <= ordertype(A,r)"; |
4091 | 244 |
by (asm_simp_tac (simpset() addsimps [ordertype_unfold, |
849 | 245 |
pred_subset RSN (2, well_ord_subset)]) 1); |
4091 | 246 |
by (fast_tac (claset() addIs [ordermap_pred_eq_ordermap] |
2925 | 247 |
addEs [predE]) 1); |
849 | 248 |
qed "ordertype_pred_subset"; |
249 |
||
5268 | 250 |
Goal "[| well_ord(A,r); x:A |] ==> \ |
849 | 251 |
\ ordertype(pred(A,x,r),r) < ordertype(A,r)"; |
252 |
by (resolve_tac [ordertype_pred_subset RS subset_imp_le RS leE] 1); |
|
253 |
by (REPEAT (ares_tac [Ord_ordertype, well_ord_subset, pred_subset] 1)); |
|
254 |
by (eresolve_tac [sym RS ordertype_eq_imp_ord_iso RS exE] 1); |
|
1461 | 255 |
by (etac well_ord_iso_predE 3); |
849 | 256 |
by (REPEAT (ares_tac [pred_subset, well_ord_subset] 1)); |
257 |
qed "ordertype_pred_lt"; |
|
258 |
||
259 |
(*May rewrite with this -- provided no rules are supplied for proving that |
|
1461 | 260 |
well_ord(pred(A,x,r), r) *) |
5268 | 261 |
Goal "well_ord(A,r) ==> \ |
849 | 262 |
\ ordertype(A,r) = {ordertype(pred(A,x,r),r). x:A}"; |
2493 | 263 |
by (rtac equalityI 1); |
4091 | 264 |
by (safe_tac (claset() addSIs [ordertype_pred_lt RS ltD])); |
984 | 265 |
by (fast_tac |
4091 | 266 |
(claset() addss |
267 |
(simpset() addsimps [ordertype_def, |
|
1461 | 268 |
well_ord_is_wf RS ordermap_eq_image, |
269 |
ordermap_type RS image_fun, |
|
270 |
ordermap_pred_eq_ordermap, |
|
271 |
pred_subset])) |
|
984 | 272 |
1); |
849 | 273 |
qed "ordertype_pred_unfold"; |
274 |
||
275 |
||
276 |
(**** Alternative definition of ordinal ****) |
|
277 |
||
278 |
(*proof by Krzysztof Grabczewski*) |
|
5137 | 279 |
Goalw [Ord_alt_def] "Ord(i) ==> Ord_alt(i)"; |
1461 | 280 |
by (rtac conjI 1); |
281 |
by (etac well_ord_Memrel 1); |
|
849 | 282 |
by (rewrite_goals_tac [Ord_def, Transset_def, pred_def, Memrel_def]); |
4091 | 283 |
by (Blast.depth_tac (claset()) 8 1); |
849 | 284 |
qed "Ord_is_Ord_alt"; |
285 |
||
286 |
(*proof by lcp*) |
|
5067 | 287 |
Goalw [Ord_alt_def, Ord_def, Transset_def, well_ord_def, |
1461 | 288 |
tot_ord_def, part_ord_def, trans_on_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
289 |
"Ord_alt(i) ==> Ord(i)"; |
9842 | 290 |
by (asm_full_simp_tac (simpset() addsimps [pred_Memrel]) 1); |
4091 | 291 |
by (blast_tac (claset() addSEs [equalityE]) 1); |
849 | 292 |
qed "Ord_alt_is_Ord"; |
293 |
||
294 |
||
295 |
(**** Ordinal Addition ****) |
|
296 |
||
297 |
(*** Order Type calculations for radd ***) |
|
298 |
||
299 |
(** Addition with 0 **) |
|
300 |
||
5067 | 301 |
Goal "(lam z:A+0. case(%x. x, %y. y, z)) : bij(A+0, A)"; |
849 | 302 |
by (res_inst_tac [("d", "Inl")] lam_bijective 1); |
4152 | 303 |
by Safe_tac; |
2925 | 304 |
by (ALLGOALS Asm_simp_tac); |
849 | 305 |
qed "bij_sum_0"; |
306 |
||
5268 | 307 |
Goal "well_ord(A,r) ==> ordertype(A+0, radd(A,r,0,s)) = ordertype(A,r)"; |
849 | 308 |
by (resolve_tac [bij_sum_0 RS ord_isoI RS ordertype_eq] 1); |
309 |
by (assume_tac 2); |
|
9842 | 310 |
by (Force_tac 1); |
849 | 311 |
qed "ordertype_sum_0_eq"; |
312 |
||
5067 | 313 |
Goal "(lam z:0+A. case(%x. x, %y. y, z)) : bij(0+A, A)"; |
849 | 314 |
by (res_inst_tac [("d", "Inr")] lam_bijective 1); |
4152 | 315 |
by Safe_tac; |
2925 | 316 |
by (ALLGOALS Asm_simp_tac); |
849 | 317 |
qed "bij_0_sum"; |
318 |
||
5268 | 319 |
Goal "well_ord(A,r) ==> ordertype(0+A, radd(0,s,A,r)) = ordertype(A,r)"; |
849 | 320 |
by (resolve_tac [bij_0_sum RS ord_isoI RS ordertype_eq] 1); |
321 |
by (assume_tac 2); |
|
9842 | 322 |
by (Force_tac 1); |
849 | 323 |
qed "ordertype_0_sum_eq"; |
324 |
||
325 |
(** Initial segments of radd. Statements by Grabczewski **) |
|
326 |
||
327 |
(*In fact, pred(A+B, Inl(a), radd(A,r,B,s)) = pred(A,a,r)+0 *) |
|
5067 | 328 |
Goalw [pred_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
329 |
"a:A ==> \ |
1461 | 330 |
\ (lam x:pred(A,a,r). Inl(x)) \ |
849 | 331 |
\ : bij(pred(A,a,r), pred(A+B, Inl(a), radd(A,r,B,s)))"; |
3840 | 332 |
by (res_inst_tac [("d", "case(%x. x, %y. y)")] lam_bijective 1); |
9842 | 333 |
by Auto_tac; |
849 | 334 |
qed "pred_Inl_bij"; |
335 |
||
5268 | 336 |
Goal "[| a:A; well_ord(A,r) |] ==> \ |
849 | 337 |
\ ordertype(pred(A+B, Inl(a), radd(A,r,B,s)), radd(A,r,B,s)) = \ |
338 |
\ ordertype(pred(A,a,r), r)"; |
|
339 |
by (resolve_tac [pred_Inl_bij RS ord_isoI RS ord_iso_sym RS ordertype_eq] 1); |
|
340 |
by (REPEAT_FIRST (ares_tac [pred_subset, well_ord_subset])); |
|
9842 | 341 |
by (asm_full_simp_tac (simpset() addsimps [pred_def]) 1); |
849 | 342 |
qed "ordertype_pred_Inl_eq"; |
343 |
||
5067 | 344 |
Goalw [pred_def, id_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
345 |
"b:B ==> \ |
1461 | 346 |
\ id(A+pred(B,b,s)) \ |
849 | 347 |
\ : bij(A+pred(B,b,s), pred(A+B, Inr(b), radd(A,r,B,s)))"; |
3840 | 348 |
by (res_inst_tac [("d", "%z. z")] lam_bijective 1); |
4152 | 349 |
by Safe_tac; |
2469 | 350 |
by (ALLGOALS (Asm_full_simp_tac)); |
849 | 351 |
qed "pred_Inr_bij"; |
352 |
||
5268 | 353 |
Goal "[| b:B; well_ord(A,r); well_ord(B,s) |] ==> \ |
849 | 354 |
\ ordertype(pred(A+B, Inr(b), radd(A,r,B,s)), radd(A,r,B,s)) = \ |
355 |
\ ordertype(A+pred(B,b,s), radd(A,r,pred(B,b,s),s))"; |
|
356 |
by (resolve_tac [pred_Inr_bij RS ord_isoI RS ord_iso_sym RS ordertype_eq] 1); |
|
4091 | 357 |
by (fast_tac (claset() addss (simpset() addsimps [pred_def, id_def])) 2); |
849 | 358 |
by (REPEAT_FIRST (ares_tac [well_ord_radd, pred_subset, well_ord_subset])); |
359 |
qed "ordertype_pred_Inr_eq"; |
|
360 |
||
361 |
(*** Basic laws for ordinal addition ***) |
|
362 |
||
5067 | 363 |
Goalw [oadd_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
364 |
"[| Ord(i); Ord(j) |] ==> Ord(i++j)"; |
849 | 365 |
by (REPEAT (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel] 1)); |
366 |
qed "Ord_oadd"; |
|
367 |
||
368 |
(** Ordinal addition with zero **) |
|
369 |
||
5137 | 370 |
Goalw [oadd_def] "Ord(i) ==> i++0 = i"; |
4091 | 371 |
by (asm_simp_tac (simpset() addsimps [Memrel_0, ordertype_sum_0_eq, |
1461 | 372 |
ordertype_Memrel, well_ord_Memrel]) 1); |
849 | 373 |
qed "oadd_0"; |
374 |
||
5137 | 375 |
Goalw [oadd_def] "Ord(i) ==> 0++i = i"; |
4091 | 376 |
by (asm_simp_tac (simpset() addsimps [Memrel_0, ordertype_0_sum_eq, |
1461 | 377 |
ordertype_Memrel, well_ord_Memrel]) 1); |
849 | 378 |
qed "oadd_0_left"; |
379 |
||
2469 | 380 |
Addsimps [oadd_0, oadd_0_left]; |
849 | 381 |
|
382 |
(*** Further properties of ordinal addition. Statements by Grabczewski, |
|
383 |
proofs by lcp. ***) |
|
384 |
||
5137 | 385 |
Goalw [oadd_def] "[| k<i; Ord(j) |] ==> k < i++j"; |
1461 | 386 |
by (rtac ltE 1 THEN assume_tac 1); |
387 |
by (rtac ltI 1); |
|
849 | 388 |
by (REPEAT (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel] 2)); |
389 |
by (asm_simp_tac |
|
4091 | 390 |
(simpset() addsimps [ordertype_pred_unfold, |
2493 | 391 |
well_ord_radd, well_ord_Memrel, |
392 |
ordertype_pred_Inl_eq, |
|
393 |
lt_pred_Memrel, leI RS le_ordertype_Memrel] |
|
2469 | 394 |
setloop rtac (InlI RSN (2,bexI))) 1); |
849 | 395 |
qed "lt_oadd1"; |
396 |
||
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
397 |
(*Thus also we obtain the rule i++j = k ==> i le k *) |
5137 | 398 |
Goal "[| Ord(i); Ord(j) |] ==> i le i++j"; |
1461 | 399 |
by (rtac all_lt_imp_le 1); |
849 | 400 |
by (REPEAT (ares_tac [Ord_oadd, lt_oadd1] 1)); |
401 |
qed "oadd_le_self"; |
|
402 |
||
403 |
(** A couple of strange but necessary results! **) |
|
404 |
||
5268 | 405 |
Goal "A<=B ==> id(A) : ord_iso(A, Memrel(A), A, Memrel(B))"; |
849 | 406 |
by (resolve_tac [id_bij RS ord_isoI] 1); |
9842 | 407 |
by (Asm_simp_tac 1); |
2925 | 408 |
by (Blast_tac 1); |
849 | 409 |
qed "id_ord_iso_Memrel"; |
410 |
||
5268 | 411 |
Goal "[| well_ord(A,r); k<j |] ==> \ |
1461 | 412 |
\ ordertype(A+k, radd(A, r, k, Memrel(j))) = \ |
849 | 413 |
\ ordertype(A+k, radd(A, r, k, Memrel(k)))"; |
1461 | 414 |
by (etac ltE 1); |
849 | 415 |
by (resolve_tac [ord_iso_refl RS sum_ord_iso_cong RS ordertype_eq] 1); |
416 |
by (eresolve_tac [OrdmemD RS id_ord_iso_Memrel RS ord_iso_sym] 1); |
|
417 |
by (REPEAT_FIRST (ares_tac [well_ord_radd, well_ord_Memrel])); |
|
418 |
qed "ordertype_sum_Memrel"; |
|
419 |
||
5137 | 420 |
Goalw [oadd_def] "[| k<j; Ord(i) |] ==> i++k < i++j"; |
1461 | 421 |
by (rtac ltE 1 THEN assume_tac 1); |
849 | 422 |
by (resolve_tac [ordertype_pred_unfold RS equalityD2 RS subsetD RS ltI] 1); |
423 |
by (REPEAT_FIRST (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel])); |
|
1461 | 424 |
by (rtac RepFun_eqI 1); |
425 |
by (etac InrI 2); |
|
849 | 426 |
by (asm_simp_tac |
4091 | 427 |
(simpset() addsimps [ordertype_pred_Inr_eq, well_ord_Memrel, |
1461 | 428 |
lt_pred_Memrel, leI RS le_ordertype_Memrel, |
429 |
ordertype_sum_Memrel]) 1); |
|
849 | 430 |
qed "oadd_lt_mono2"; |
431 |
||
5268 | 432 |
Goal "[| i++j < i++k; Ord(i); Ord(j); Ord(k) |] ==> j<k"; |
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
433 |
by (rtac Ord_linear_lt 1); |
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
434 |
by (REPEAT_SOME assume_tac); |
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
435 |
by (ALLGOALS |
4091 | 436 |
(blast_tac (claset() addDs [oadd_lt_mono2] addEs [lt_irrefl, lt_asym]))); |
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
437 |
qed "oadd_lt_cancel2"; |
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
438 |
|
5268 | 439 |
Goal "[| Ord(i); Ord(j); Ord(k) |] ==> i++j < i++k <-> j<k"; |
4091 | 440 |
by (blast_tac (claset() addSIs [oadd_lt_mono2] addSDs [oadd_lt_cancel2]) 1); |
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
441 |
qed "oadd_lt_iff2"; |
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
442 |
|
5268 | 443 |
Goal "[| i++j = i++k; Ord(i); Ord(j); Ord(k) |] ==> j=k"; |
849 | 444 |
by (rtac Ord_linear_lt 1); |
445 |
by (REPEAT_SOME assume_tac); |
|
446 |
by (ALLGOALS |
|
4091 | 447 |
(fast_tac (claset() addDs [oadd_lt_mono2] |
448 |
addss (simpset() addsimps [lt_not_refl])))); |
|
849 | 449 |
qed "oadd_inject"; |
450 |
||
5067 | 451 |
Goalw [oadd_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
452 |
"[| k < i++j; Ord(i); Ord(j) |] ==> k<i | (EX l:j. k = i++l )"; |
849 | 453 |
(*Rotate the hypotheses so that simplification will work*) |
454 |
by (etac revcut_rl 1); |
|
455 |
by (asm_full_simp_tac |
|
4091 | 456 |
(simpset() addsimps [ordertype_pred_unfold, well_ord_radd, |
1461 | 457 |
well_ord_Memrel]) 1); |
849 | 458 |
by (eresolve_tac [ltD RS RepFunE] 1); |
4091 | 459 |
by (fast_tac (claset() addss |
460 |
(simpset() addsimps [ordertype_pred_Inl_eq, well_ord_Memrel, |
|
1461 | 461 |
ltI, lt_pred_Memrel, le_ordertype_Memrel, leI, |
462 |
ordertype_pred_Inr_eq, |
|
463 |
ordertype_sum_Memrel])) 1); |
|
849 | 464 |
qed "lt_oadd_disj"; |
465 |
||
466 |
||
467 |
(*** Ordinal addition with successor -- via associativity! ***) |
|
468 |
||
5067 | 469 |
Goalw [oadd_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
470 |
"[| Ord(i); Ord(j); Ord(k) |] ==> (i++j)++k = i++(j++k)"; |
849 | 471 |
by (resolve_tac [ordertype_eq RS trans] 1); |
472 |
by (rtac ([ordertype_ord_iso RS ord_iso_sym, ord_iso_refl] MRS |
|
1461 | 473 |
sum_ord_iso_cong) 1); |
849 | 474 |
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel, Ord_ordertype] 1)); |
475 |
by (resolve_tac [sum_assoc_ord_iso RS ordertype_eq RS trans] 1); |
|
476 |
by (rtac ([ord_iso_refl, ordertype_ord_iso] MRS sum_ord_iso_cong RS |
|
1461 | 477 |
ordertype_eq) 2); |
849 | 478 |
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel, Ord_ordertype] 1)); |
479 |
qed "oadd_assoc"; |
|
480 |
||
5268 | 481 |
Goal "[| Ord(i); Ord(j) |] ==> i++j = i Un (UN k:j. {i++k})"; |
849 | 482 |
by (rtac (subsetI RS equalityI) 1); |
483 |
by (eresolve_tac [ltI RS lt_oadd_disj RS disjE] 1); |
|
484 |
by (REPEAT (ares_tac [Ord_oadd] 1)); |
|
4091 | 485 |
by (fast_tac (claset() addIs [lt_oadd1, oadd_lt_mono2] |
486 |
addss (simpset() addsimps [Ord_mem_iff_lt, Ord_oadd])) 3); |
|
2925 | 487 |
by (Blast_tac 2); |
4091 | 488 |
by (blast_tac (claset() addSEs [ltE]) 1); |
849 | 489 |
qed "oadd_unfold"; |
490 |
||
5137 | 491 |
Goal "Ord(i) ==> i++1 = succ(i)"; |
4091 | 492 |
by (asm_simp_tac (simpset() addsimps [oadd_unfold, Ord_1, oadd_0]) 1); |
2925 | 493 |
by (Blast_tac 1); |
849 | 494 |
qed "oadd_1"; |
495 |
||
5268 | 496 |
Goal "[| Ord(i); Ord(j) |] ==> i++succ(j) = succ(i++j)"; |
6153 | 497 |
(*FOL_ss prevents looping*) |
498 |
by (asm_simp_tac (FOL_ss addsimps [Ord_oadd, oadd_1 RS sym]) 1); |
|
4091 | 499 |
by (asm_simp_tac (simpset() addsimps [oadd_1, oadd_assoc, Ord_1]) 1); |
849 | 500 |
qed "oadd_succ"; |
501 |
||
502 |
||
503 |
(** Ordinal addition with limit ordinals **) |
|
504 |
||
5268 | 505 |
val prems = |
506 |
Goal "[| Ord(i); !!x. x:A ==> Ord(j(x)); a:A |] ==> \ |
|
507 |
\ i ++ (UN x:A. j(x)) = (UN x:A. i++j(x))"; |
|
5618 | 508 |
by (blast_tac (claset() addIs prems @ [ltI, Ord_UN, Ord_oadd, |
509 |
lt_oadd1 RS ltD, oadd_lt_mono2 RS ltD] |
|
9173 | 510 |
addSEs [ltE] addSDs [ltI RS lt_oadd_disj]) 1); |
849 | 511 |
qed "oadd_UN"; |
512 |
||
5268 | 513 |
Goal "[| Ord(i); Limit(j) |] ==> i++j = (UN k:j. i++k)"; |
849 | 514 |
by (forward_tac [Limit_has_0 RS ltD] 1); |
4091 | 515 |
by (asm_simp_tac (simpset() addsimps [Limit_is_Ord RS Ord_in_Ord, |
9173 | 516 |
oadd_UN RS sym, Union_eq_UN RS sym, |
517 |
Limit_Union_eq]) 1); |
|
849 | 518 |
qed "oadd_Limit"; |
519 |
||
520 |
(** Order/monotonicity properties of ordinal addition **) |
|
521 |
||
5137 | 522 |
Goal "[| Ord(i); Ord(j) |] ==> i le j++i"; |
849 | 523 |
by (eres_inst_tac [("i","i")] trans_induct3 1); |
4091 | 524 |
by (asm_simp_tac (simpset() addsimps [Ord_0_le]) 1); |
525 |
by (asm_simp_tac (simpset() addsimps [oadd_succ, succ_leI]) 1); |
|
526 |
by (asm_simp_tac (simpset() addsimps [oadd_Limit]) 1); |
|
1461 | 527 |
by (rtac le_trans 1); |
528 |
by (rtac le_implies_UN_le_UN 2); |
|
2925 | 529 |
by (Blast_tac 2); |
4091 | 530 |
by (asm_simp_tac (simpset() addsimps [Union_eq_UN RS sym, Limit_Union_eq, |
2493 | 531 |
le_refl, Limit_is_Ord]) 1); |
849 | 532 |
qed "oadd_le_self2"; |
533 |
||
5137 | 534 |
Goal "[| k le j; Ord(i) |] ==> k++i le j++i"; |
7499 | 535 |
by (ftac lt_Ord 1); |
536 |
by (ftac le_Ord2 1); |
|
1461 | 537 |
by (etac trans_induct3 1); |
2469 | 538 |
by (Asm_simp_tac 1); |
4091 | 539 |
by (asm_simp_tac (simpset() addsimps [oadd_succ, succ_le_iff]) 1); |
540 |
by (asm_simp_tac (simpset() addsimps [oadd_Limit]) 1); |
|
1461 | 541 |
by (rtac le_implies_UN_le_UN 1); |
2925 | 542 |
by (Blast_tac 1); |
849 | 543 |
qed "oadd_le_mono1"; |
544 |
||
5137 | 545 |
Goal "[| i' le i; j'<j |] ==> i'++j' < i++j"; |
1461 | 546 |
by (rtac lt_trans1 1); |
849 | 547 |
by (REPEAT (eresolve_tac [asm_rl, oadd_le_mono1, oadd_lt_mono2, ltE, |
1461 | 548 |
Ord_succD] 1)); |
849 | 549 |
qed "oadd_lt_mono"; |
550 |
||
5137 | 551 |
Goal "[| i' le i; j' le j |] ==> i'++j' le i++j"; |
4091 | 552 |
by (asm_simp_tac (simpset() addsimps [oadd_succ RS sym, le_Ord2, oadd_lt_mono]) 1); |
849 | 553 |
qed "oadd_le_mono"; |
554 |
||
5268 | 555 |
Goal "[| Ord(i); Ord(j); Ord(k) |] ==> i++j le i++k <-> j le k"; |
4091 | 556 |
by (asm_simp_tac (simpset() addsimps [oadd_lt_iff2, oadd_succ RS sym, |
1461 | 557 |
Ord_succ]) 1); |
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
558 |
qed "oadd_le_iff2"; |
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
559 |
|
849 | 560 |
|
561 |
(** Ordinal subtraction; the difference is ordertype(j-i, Memrel(j)). |
|
562 |
Probably simpler to define the difference recursively! |
|
563 |
**) |
|
564 |
||
5268 | 565 |
Goal "A<=B ==> (lam y:B. if(y:A, Inl(y), Inr(y))) : bij(B, A+(B-A))"; |
3840 | 566 |
by (res_inst_tac [("d", "case(%x. x, %y. y)")] lam_bijective 1); |
4091 | 567 |
by (blast_tac (claset() addSIs [if_type]) 1); |
568 |
by (fast_tac (claset() addSIs [case_type]) 1); |
|
1461 | 569 |
by (etac sumE 2); |
5137 | 570 |
by (ALLGOALS Asm_simp_tac); |
849 | 571 |
qed "bij_sum_Diff"; |
572 |
||
5268 | 573 |
Goal "i le j ==> \ |
1461 | 574 |
\ ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j))) = \ |
849 | 575 |
\ ordertype(j, Memrel(j))"; |
4091 | 576 |
by (safe_tac (claset() addSDs [le_subset_iff RS iffD1])); |
849 | 577 |
by (resolve_tac [bij_sum_Diff RS ord_isoI RS ord_iso_sym RS ordertype_eq] 1); |
1461 | 578 |
by (etac well_ord_Memrel 3); |
849 | 579 |
by (assume_tac 1); |
9842 | 580 |
by (Asm_simp_tac 1); |
849 | 581 |
by (forw_inst_tac [("j", "y")] Ord_in_Ord 1 THEN assume_tac 1); |
582 |
by (forw_inst_tac [("j", "x")] Ord_in_Ord 1 THEN assume_tac 1); |
|
4091 | 583 |
by (asm_simp_tac (simpset() addsimps [Ord_mem_iff_lt, lt_Ord, not_lt_iff_le]) 1); |
584 |
by (blast_tac (claset() addIs [lt_trans2, lt_trans]) 1); |
|
849 | 585 |
qed "ordertype_sum_Diff"; |
586 |
||
5067 | 587 |
Goalw [oadd_def, odiff_def] |
5268 | 588 |
"i le j \ |
589 |
\ ==> i ++ (j--i) = ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j)))"; |
|
4091 | 590 |
by (safe_tac (claset() addSDs [le_subset_iff RS iffD1])); |
849 | 591 |
by (resolve_tac [sum_ord_iso_cong RS ordertype_eq] 1); |
1461 | 592 |
by (etac id_ord_iso_Memrel 1); |
849 | 593 |
by (resolve_tac [ordertype_ord_iso RS ord_iso_sym] 1); |
594 |
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel RS well_ord_subset, |
|
1461 | 595 |
Diff_subset] 1)); |
849 | 596 |
qed "oadd_ordertype_Diff"; |
597 |
||
5137 | 598 |
Goal "i le j ==> i ++ (j--i) = j"; |
4091 | 599 |
by (asm_simp_tac (simpset() addsimps [oadd_ordertype_Diff, ordertype_sum_Diff, |
1461 | 600 |
ordertype_Memrel, lt_Ord2 RS Ord_succD]) 1); |
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
601 |
qed "oadd_odiff_inverse"; |
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
602 |
|
5067 | 603 |
Goalw [odiff_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
604 |
"[| Ord(i); Ord(j) |] ==> Ord(i--j)"; |
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
605 |
by (REPEAT (ares_tac [Ord_ordertype, well_ord_Memrel RS well_ord_subset, |
1461 | 606 |
Diff_subset] 1)); |
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
607 |
qed "Ord_odiff"; |
849 | 608 |
|
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
609 |
(*By oadd_inject, the difference between i and j is unique. Note that we get |
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
610 |
i++j = k ==> j = k--i. *) |
5268 | 611 |
Goal "[| Ord(i); Ord(j) |] ==> (i++j) -- i = j"; |
1461 | 612 |
by (rtac oadd_inject 1); |
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
613 |
by (REPEAT (ares_tac [Ord_ordertype, Ord_oadd, Ord_odiff] 2)); |
4091 | 614 |
by (asm_simp_tac (simpset() addsimps [oadd_odiff_inverse, oadd_le_self]) 1); |
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
615 |
qed "odiff_oadd_inverse"; |
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
616 |
|
9907 | 617 |
val [i_lt_j, k_le_i] = goal (the_context ()) |
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
618 |
"[| i<j; k le i |] ==> i--k < j--k"; |
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
619 |
by (rtac (k_le_i RS lt_Ord RSN (2,oadd_lt_cancel2)) 1); |
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
620 |
by (simp_tac |
4091 | 621 |
(simpset() addsimps [i_lt_j, k_le_i, [k_le_i, leI] MRS le_trans, |
5268 | 622 |
oadd_odiff_inverse]) 1); |
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
623 |
by (REPEAT (resolve_tac (Ord_odiff :: |
1461 | 624 |
([i_lt_j, k_le_i] RL [lt_Ord, lt_Ord2])) 1)); |
1032
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
lcp
parents:
984
diff
changeset
|
625 |
qed "odiff_lt_mono2"; |
849 | 626 |
|
627 |
||
628 |
(**** Ordinal Multiplication ****) |
|
629 |
||
5067 | 630 |
Goalw [omult_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
631 |
"[| Ord(i); Ord(j) |] ==> Ord(i**j)"; |
849 | 632 |
by (REPEAT (ares_tac [Ord_ordertype, well_ord_rmult, well_ord_Memrel] 1)); |
633 |
qed "Ord_omult"; |
|
634 |
||
635 |
(*** A useful unfolding law ***) |
|
636 |
||
5067 | 637 |
Goalw [pred_def] |
9842 | 638 |
"[| a:A; b:B |] ==> pred(A*B, <a,b>, rmult(A,r,B,s)) = \ |
639 |
\ pred(A,a,r)*B Un ({a} * pred(B,b,s))"; |
|
2493 | 640 |
by (rtac equalityI 1); |
4152 | 641 |
by Safe_tac; |
9842 | 642 |
by (ALLGOALS Asm_full_simp_tac); |
643 |
by (ALLGOALS Blast_tac); |
|
849 | 644 |
qed "pred_Pair_eq"; |
645 |
||
5268 | 646 |
Goal "[| a:A; b:B; well_ord(A,r); well_ord(B,s) |] ==> \ |
849 | 647 |
\ ordertype(pred(A*B, <a,b>, rmult(A,r,B,s)), rmult(A,r,B,s)) = \ |
3016 | 648 |
\ ordertype(pred(A,a,r)*B + pred(B,b,s), \ |
849 | 649 |
\ radd(A*B, rmult(A,r,B,s), B, s))"; |
4091 | 650 |
by (asm_simp_tac (simpset() addsimps [pred_Pair_eq]) 1); |
849 | 651 |
by (resolve_tac [ordertype_eq RS sym] 1); |
1461 | 652 |
by (rtac prod_sum_singleton_ord_iso 1); |
984 | 653 |
by (REPEAT_FIRST (ares_tac [pred_subset, well_ord_rmult RS well_ord_subset])); |
4091 | 654 |
by (blast_tac (claset() addSEs [predE]) 1); |
849 | 655 |
qed "ordertype_pred_Pair_eq"; |
656 |
||
5067 | 657 |
Goalw [oadd_def, omult_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
658 |
"[| i'<i; j'<j |] ==> \ |
849 | 659 |
\ ordertype(pred(i*j, <i',j'>, rmult(i,Memrel(i),j,Memrel(j))), \ |
3016 | 660 |
\ rmult(i,Memrel(i),j,Memrel(j))) = \ |
849 | 661 |
\ j**i' ++ j'"; |
4091 | 662 |
by (asm_simp_tac (simpset() addsimps [ordertype_pred_Pair_eq, lt_pred_Memrel, |
3016 | 663 |
ltD, lt_Ord2, well_ord_Memrel]) 1); |
1461 | 664 |
by (rtac trans 1); |
849 | 665 |
by (resolve_tac [ordertype_ord_iso RS sum_ord_iso_cong RS ordertype_eq] 2); |
1461 | 666 |
by (rtac ord_iso_refl 3); |
849 | 667 |
by (resolve_tac [id_bij RS ord_isoI RS ordertype_eq] 1); |
668 |
by (REPEAT_FIRST (eresolve_tac [SigmaE, sumE, ltE, ssubst])); |
|
669 |
by (REPEAT_FIRST (ares_tac [well_ord_rmult, well_ord_radd, well_ord_Memrel, |
|
1461 | 670 |
Ord_ordertype])); |
9842 | 671 |
by (ALLGOALS Asm_simp_tac); |
4152 | 672 |
by Safe_tac; |
4091 | 673 |
by (ALLGOALS (blast_tac (claset() addIs [Ord_trans]))); |
849 | 674 |
qed "ordertype_pred_Pair_lemma"; |
675 |
||
5067 | 676 |
Goalw [omult_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
677 |
"[| Ord(i); Ord(j); k<j**i |] ==> \ |
849 | 678 |
\ EX j' i'. k = j**i' ++ j' & j'<j & i'<i"; |
4091 | 679 |
by (asm_full_simp_tac (simpset() addsimps [ordertype_pred_unfold, |
1461 | 680 |
well_ord_rmult, well_ord_Memrel]) 1); |
4091 | 681 |
by (safe_tac (claset() addSEs [ltE])); |
682 |
by (asm_simp_tac (simpset() addsimps [ordertype_pred_Pair_lemma, ltI, |
|
3736
39ee3d31cfbc
Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents:
3016
diff
changeset
|
683 |
symmetric omult_def]) 1); |
4091 | 684 |
by (blast_tac (claset() addIs [ltI]) 1); |
849 | 685 |
qed "lt_omult"; |
686 |
||
5067 | 687 |
Goalw [omult_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
688 |
"[| j'<j; i'<i |] ==> j**i' ++ j' < j**i"; |
1461 | 689 |
by (rtac ltI 1); |
984 | 690 |
by (asm_simp_tac |
4091 | 691 |
(simpset() addsimps [Ord_ordertype, well_ord_rmult, well_ord_Memrel, |
2493 | 692 |
lt_Ord2]) 2); |
984 | 693 |
by (asm_simp_tac |
4091 | 694 |
(simpset() addsimps [ordertype_pred_unfold, |
1461 | 695 |
well_ord_rmult, well_ord_Memrel, lt_Ord2]) 1); |
2469 | 696 |
by (rtac bexI 1); |
4091 | 697 |
by (blast_tac (claset() addSEs [ltE]) 2); |
849 | 698 |
by (asm_simp_tac |
4091 | 699 |
(simpset() addsimps [ordertype_pred_Pair_lemma, ltI, |
2493 | 700 |
symmetric omult_def]) 1); |
849 | 701 |
qed "omult_oadd_lt"; |
702 |
||
5268 | 703 |
Goal "[| Ord(i); Ord(j) |] ==> j**i = (UN j':j. UN i':i. {j**i' ++ j'})"; |
849 | 704 |
by (rtac (subsetI RS equalityI) 1); |
705 |
by (resolve_tac [lt_omult RS exE] 1); |
|
1461 | 706 |
by (etac ltI 3); |
849 | 707 |
by (REPEAT (ares_tac [Ord_omult] 1)); |
4091 | 708 |
by (blast_tac (claset() addSEs [ltE]) 1); |
709 |
by (blast_tac (claset() addIs [omult_oadd_lt RS ltD, ltI]) 1); |
|
849 | 710 |
qed "omult_unfold"; |
711 |
||
712 |
(*** Basic laws for ordinal multiplication ***) |
|
713 |
||
714 |
(** Ordinal multiplication by zero **) |
|
715 |
||
5067 | 716 |
Goalw [omult_def] "i**0 = 0"; |
2469 | 717 |
by (Asm_simp_tac 1); |
849 | 718 |
qed "omult_0"; |
719 |
||
5067 | 720 |
Goalw [omult_def] "0**i = 0"; |
2469 | 721 |
by (Asm_simp_tac 1); |
849 | 722 |
qed "omult_0_left"; |
723 |
||
2469 | 724 |
Addsimps [omult_0, omult_0_left]; |
725 |
||
849 | 726 |
(** Ordinal multiplication by 1 **) |
727 |
||
5137 | 728 |
Goalw [omult_def] "Ord(i) ==> i**1 = i"; |
849 | 729 |
by (resolve_tac [ord_isoI RS ordertype_eq RS trans] 1); |
730 |
by (res_inst_tac [("c", "snd"), ("d", "%z.<0,z>")] lam_bijective 1); |
|
731 |
by (REPEAT_FIRST (eresolve_tac [snd_type, SigmaE, succE, emptyE, |
|
1461 | 732 |
well_ord_Memrel, ordertype_Memrel])); |
9842 | 733 |
by (ALLGOALS Asm_simp_tac); |
849 | 734 |
qed "omult_1"; |
735 |
||
5137 | 736 |
Goalw [omult_def] "Ord(i) ==> 1**i = i"; |
849 | 737 |
by (resolve_tac [ord_isoI RS ordertype_eq RS trans] 1); |
738 |
by (res_inst_tac [("c", "fst"), ("d", "%z.<z,0>")] lam_bijective 1); |
|
739 |
by (REPEAT_FIRST (eresolve_tac [fst_type, SigmaE, succE, emptyE, |
|
1461 | 740 |
well_ord_Memrel, ordertype_Memrel])); |
9842 | 741 |
by (ALLGOALS Asm_simp_tac); |
849 | 742 |
qed "omult_1_left"; |
743 |
||
2469 | 744 |
Addsimps [omult_1, omult_1_left]; |
745 |
||
849 | 746 |
(** Distributive law for ordinal multiplication and addition **) |
747 |
||
5067 | 748 |
Goalw [omult_def, oadd_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
749 |
"[| Ord(i); Ord(j); Ord(k) |] ==> i**(j++k) = (i**j)++(i**k)"; |
849 | 750 |
by (resolve_tac [ordertype_eq RS trans] 1); |
751 |
by (rtac ([ordertype_ord_iso RS ord_iso_sym, ord_iso_refl] MRS |
|
1461 | 752 |
prod_ord_iso_cong) 1); |
849 | 753 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_radd, well_ord_Memrel, |
1461 | 754 |
Ord_ordertype] 1)); |
849 | 755 |
by (rtac (sum_prod_distrib_ord_iso RS ordertype_eq RS trans) 1); |
756 |
by (rtac ordertype_eq 2); |
|
757 |
by (rtac ([ordertype_ord_iso, ordertype_ord_iso] MRS sum_ord_iso_cong) 2); |
|
758 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_radd, well_ord_Memrel, |
|
1461 | 759 |
Ord_ordertype] 1)); |
849 | 760 |
qed "oadd_omult_distrib"; |
761 |
||
5137 | 762 |
Goal "[| Ord(i); Ord(j) |] ==> i**succ(j) = (i**j)++i"; |
6153 | 763 |
(*FOL_ss prevents looping*) |
764 |
by (asm_simp_tac (FOL_ss addsimps [oadd_1 RS sym]) 1); |
|
849 | 765 |
by (asm_simp_tac |
4091 | 766 |
(simpset() addsimps [omult_1, oadd_omult_distrib, Ord_1]) 1); |
849 | 767 |
qed "omult_succ"; |
768 |
||
769 |
(** Associative law **) |
|
770 |
||
5067 | 771 |
Goalw [omult_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
772 |
"[| Ord(i); Ord(j); Ord(k) |] ==> (i**j)**k = i**(j**k)"; |
849 | 773 |
by (resolve_tac [ordertype_eq RS trans] 1); |
774 |
by (rtac ([ord_iso_refl, ordertype_ord_iso RS ord_iso_sym] MRS |
|
1461 | 775 |
prod_ord_iso_cong) 1); |
849 | 776 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1)); |
777 |
by (resolve_tac [prod_assoc_ord_iso RS ord_iso_sym RS |
|
1461 | 778 |
ordertype_eq RS trans] 1); |
849 | 779 |
by (rtac ([ordertype_ord_iso, ord_iso_refl] MRS prod_ord_iso_cong RS |
1461 | 780 |
ordertype_eq) 2); |
849 | 781 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel, Ord_ordertype] 1)); |
782 |
qed "omult_assoc"; |
|
783 |
||
784 |
||
785 |
(** Ordinal multiplication with limit ordinals **) |
|
786 |
||
5268 | 787 |
val prems = |
788 |
Goal "[| Ord(i); !!x. x:A ==> Ord(j(x)) |] ==> \ |
|
789 |
\ i ** (UN x:A. j(x)) = (UN x:A. i**j(x))"; |
|
5529 | 790 |
by (asm_simp_tac (simpset() addsimps prems @ [Ord_UN, omult_unfold]) 1); |
2925 | 791 |
by (Blast_tac 1); |
849 | 792 |
qed "omult_UN"; |
467 | 793 |
|
5268 | 794 |
Goal "[| Ord(i); Limit(j) |] ==> i**j = (UN k:j. i**k)"; |
849 | 795 |
by (asm_simp_tac |
4091 | 796 |
(simpset() addsimps [Limit_is_Ord RS Ord_in_Ord, omult_UN RS sym, |
1461 | 797 |
Union_eq_UN RS sym, Limit_Union_eq]) 1); |
849 | 798 |
qed "omult_Limit"; |
799 |
||
800 |
||
801 |
(*** Ordering/monotonicity properties of ordinal multiplication ***) |
|
802 |
||
803 |
(*As a special case we have "[| 0<i; 0<j |] ==> 0 < i**j" *) |
|
5137 | 804 |
Goal "[| k<i; 0<j |] ==> k < i**j"; |
4091 | 805 |
by (safe_tac (claset() addSEs [ltE] addSIs [ltI, Ord_omult])); |
806 |
by (asm_simp_tac (simpset() addsimps [omult_unfold]) 1); |
|
2469 | 807 |
by (REPEAT_FIRST (ares_tac [bexI])); |
808 |
by (Asm_simp_tac 1); |
|
849 | 809 |
qed "lt_omult1"; |
810 |
||
5137 | 811 |
Goal "[| Ord(i); 0<j |] ==> i le i**j"; |
1461 | 812 |
by (rtac all_lt_imp_le 1); |
849 | 813 |
by (REPEAT (ares_tac [Ord_omult, lt_omult1, lt_Ord2] 1)); |
814 |
qed "omult_le_self"; |
|
815 |
||
5137 | 816 |
Goal "[| k le j; Ord(i) |] ==> k**i le j**i"; |
7499 | 817 |
by (ftac lt_Ord 1); |
818 |
by (ftac le_Ord2 1); |
|
1461 | 819 |
by (etac trans_induct3 1); |
4091 | 820 |
by (asm_simp_tac (simpset() addsimps [le_refl, Ord_0]) 1); |
821 |
by (asm_simp_tac (simpset() addsimps [omult_succ, oadd_le_mono]) 1); |
|
822 |
by (asm_simp_tac (simpset() addsimps [omult_Limit]) 1); |
|
1461 | 823 |
by (rtac le_implies_UN_le_UN 1); |
2925 | 824 |
by (Blast_tac 1); |
849 | 825 |
qed "omult_le_mono1"; |
826 |
||
5137 | 827 |
Goal "[| k<j; 0<i |] ==> i**k < i**j"; |
1461 | 828 |
by (rtac ltI 1); |
4091 | 829 |
by (asm_simp_tac (simpset() addsimps [omult_unfold, lt_Ord2]) 1); |
830 |
by (safe_tac (claset() addSEs [ltE] addSIs [Ord_omult])); |
|
2469 | 831 |
by (REPEAT_FIRST (ares_tac [bexI])); |
4091 | 832 |
by (asm_simp_tac (simpset() addsimps [Ord_omult]) 1); |
849 | 833 |
qed "omult_lt_mono2"; |
834 |
||
5137 | 835 |
Goal "[| k le j; Ord(i) |] ==> i**k le i**j"; |
1461 | 836 |
by (rtac subset_imp_le 1); |
4091 | 837 |
by (safe_tac (claset() addSEs [ltE, make_elim Ord_succD] addSIs [Ord_omult])); |
838 |
by (asm_full_simp_tac (simpset() addsimps [omult_unfold]) 1); |
|
839 |
by (deepen_tac (claset() addEs [Ord_trans]) 0 1); |
|
849 | 840 |
qed "omult_le_mono2"; |
841 |
||
5137 | 842 |
Goal "[| i' le i; j' le j |] ==> i'**j' le i**j"; |
1461 | 843 |
by (rtac le_trans 1); |
849 | 844 |
by (REPEAT (eresolve_tac [asm_rl, omult_le_mono1, omult_le_mono2, ltE, |
1461 | 845 |
Ord_succD] 1)); |
849 | 846 |
qed "omult_le_mono"; |
847 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
848 |
Goal "[| i' le i; j'<j; 0<i |] ==> i'**j' < i**j"; |
1461 | 849 |
by (rtac lt_trans1 1); |
849 | 850 |
by (REPEAT (eresolve_tac [asm_rl, omult_le_mono1, omult_lt_mono2, ltE, |
1461 | 851 |
Ord_succD] 1)); |
849 | 852 |
qed "omult_lt_mono"; |
853 |
||
5137 | 854 |
Goal "[| Ord(i); 0<j |] ==> i le j**i"; |
7499 | 855 |
by (ftac lt_Ord2 1); |
849 | 856 |
by (eres_inst_tac [("i","i")] trans_induct3 1); |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
857 |
by (Asm_simp_tac 1); |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
858 |
by (asm_simp_tac (simpset() addsimps [omult_succ]) 1); |
1461 | 859 |
by (etac lt_trans1 1); |
849 | 860 |
by (res_inst_tac [("b", "j**x")] (oadd_0 RS subst) 1 THEN |
861 |
rtac oadd_lt_mono2 2); |
|
862 |
by (REPEAT (ares_tac [Ord_omult] 1)); |
|
4091 | 863 |
by (asm_simp_tac (simpset() addsimps [omult_Limit]) 1); |
1461 | 864 |
by (rtac le_trans 1); |
865 |
by (rtac le_implies_UN_le_UN 2); |
|
2925 | 866 |
by (Blast_tac 2); |
4091 | 867 |
by (asm_simp_tac (simpset() addsimps [Union_eq_UN RS sym, Limit_Union_eq, |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
868 |
Limit_is_Ord]) 1); |
849 | 869 |
qed "omult_le_self2"; |
870 |
||
871 |
||
872 |
(** Further properties of ordinal multiplication **) |
|
873 |
||
5137 | 874 |
Goal "[| i**j = i**k; 0<i; Ord(j); Ord(k) |] ==> j=k"; |
849 | 875 |
by (rtac Ord_linear_lt 1); |
876 |
by (REPEAT_SOME assume_tac); |
|
877 |
by (ALLGOALS |
|
4091 | 878 |
(best_tac (claset() addDs [omult_lt_mono2] |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
879 |
addss (simpset() addsimps [lt_not_refl])))); |
849 | 880 |
qed "omult_inject"; |
881 |
||
882 |