| author | wenzelm | 
| Wed, 09 Aug 2000 20:46:58 +0200 | |
| changeset 9562 | 6b07b56aa3a8 | 
| parent 9173 | 422968aeed49 | 
| child 9842 | 58d8335cc40c | 
| 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  | 
|
| 849 | 15  | 
val [prem] = goal OrderType.thy "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);  | 
| 4091 | 19  | 
by (asm_simp_tac (simpset() addsimps [linear_def, Memrel_iff,  | 
| 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";  | 
| 4091 | 32  | 
by (asm_simp_tac (simpset() addsimps [Memrel_iff]) 1);  | 
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);  | 
| 4091 | 46  | 
by (asm_full_simp_tac (simpset() addsimps [ord_iso_def]) 1);  | 
| 
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*)  | 
| 4091 | 48  | 
by (asm_simp_tac (simpset() addsimps [Memrel_iff]) 1);  | 
49  | 
by (fast_tac (claset() addSEs [bij_is_fun RS apply_type] addEs [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*)  | 
| 5268 | 53  | 
Goal "[| Ord(i); Ord(j); f: ord_iso(i,Memrel(i),j,Memrel(j)) \  | 
| 
814
 
a32b420c33d4
Moved well_ord_Memrel, lt_eq_pred, Ord_iso_implies_eq_lemma,
 
lcp 
parents: 
807 
diff
changeset
 | 
54  | 
\ |] ==> i=j";  | 
| 
 
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!*)  | 
|
90  | 
val ordermap_unfold = rewrite_rule [pred_def] ordermap_pred_unfold;  | 
|
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);  | 
|
| 4091 | 197  | 
by (asm_simp_tac (simpset() addsimps [id_conv, Memrel_iff]) 1);  | 
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)";  | 
| 4091 | 290  | 
by (asm_full_simp_tac (simpset() addsimps [Memrel_iff, pred_Memrel]) 1);  | 
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);  | 
|
| 4091 | 310  | 
by (fast_tac (claset() addss (simpset() addsimps [radd_Inl_iff, Memrel_iff])) 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);  | 
|
| 4091 | 322  | 
by (fast_tac (claset() addss (simpset() addsimps [radd_Inr_iff, Memrel_iff])) 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);
 | 
| 4152 | 333  | 
by Safe_tac;  | 
| 849 | 334  | 
by (ALLGOALS  | 
335  | 
(asm_full_simp_tac  | 
|
| 4091 | 336  | 
(simpset() addsimps [radd_Inl_iff, radd_Inr_Inl_iff])));  | 
| 849 | 337  | 
qed "pred_Inl_bij";  | 
338  | 
||
| 5268 | 339  | 
Goal "[| a:A; well_ord(A,r) |] ==> \  | 
| 849 | 340  | 
\ ordertype(pred(A+B, Inl(a), radd(A,r,B,s)), radd(A,r,B,s)) = \  | 
341  | 
\ ordertype(pred(A,a,r), r)";  | 
|
342  | 
by (resolve_tac [pred_Inl_bij RS ord_isoI RS ord_iso_sym RS ordertype_eq] 1);  | 
|
343  | 
by (REPEAT_FIRST (ares_tac [pred_subset, well_ord_subset]));  | 
|
| 4091 | 344  | 
by (asm_full_simp_tac (simpset() addsimps [radd_Inl_iff, pred_def]) 1);  | 
| 849 | 345  | 
qed "ordertype_pred_Inl_eq";  | 
346  | 
||
| 5067 | 347  | 
Goalw [pred_def, id_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
348  | 
"b:B ==> \  | 
| 1461 | 349  | 
\ id(A+pred(B,b,s)) \  | 
| 849 | 350  | 
\ : bij(A+pred(B,b,s), pred(A+B, Inr(b), radd(A,r,B,s)))";  | 
| 3840 | 351  | 
by (res_inst_tac [("d", "%z. z")] lam_bijective 1);
 | 
| 4152 | 352  | 
by Safe_tac;  | 
| 2469 | 353  | 
by (ALLGOALS (Asm_full_simp_tac));  | 
| 849 | 354  | 
qed "pred_Inr_bij";  | 
355  | 
||
| 5268 | 356  | 
Goal "[| b:B; well_ord(A,r); well_ord(B,s) |] ==> \  | 
| 849 | 357  | 
\ ordertype(pred(A+B, Inr(b), radd(A,r,B,s)), radd(A,r,B,s)) = \  | 
358  | 
\ ordertype(A+pred(B,b,s), radd(A,r,pred(B,b,s),s))";  | 
|
359  | 
by (resolve_tac [pred_Inr_bij RS ord_isoI RS ord_iso_sym RS ordertype_eq] 1);  | 
|
| 4091 | 360  | 
by (fast_tac (claset() addss (simpset() addsimps [pred_def, id_def])) 2);  | 
| 849 | 361  | 
by (REPEAT_FIRST (ares_tac [well_ord_radd, pred_subset, well_ord_subset]));  | 
362  | 
qed "ordertype_pred_Inr_eq";  | 
|
363  | 
||
364  | 
(*** Basic laws for ordinal addition ***)  | 
|
365  | 
||
| 5067 | 366  | 
Goalw [oadd_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
367  | 
"[| Ord(i); Ord(j) |] ==> Ord(i++j)";  | 
| 849 | 368  | 
by (REPEAT (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel] 1));  | 
369  | 
qed "Ord_oadd";  | 
|
370  | 
||
371  | 
(** Ordinal addition with zero **)  | 
|
372  | 
||
| 5137 | 373  | 
Goalw [oadd_def] "Ord(i) ==> i++0 = i";  | 
| 4091 | 374  | 
by (asm_simp_tac (simpset() addsimps [Memrel_0, ordertype_sum_0_eq,  | 
| 1461 | 375  | 
ordertype_Memrel, well_ord_Memrel]) 1);  | 
| 849 | 376  | 
qed "oadd_0";  | 
377  | 
||
| 5137 | 378  | 
Goalw [oadd_def] "Ord(i) ==> 0++i = i";  | 
| 4091 | 379  | 
by (asm_simp_tac (simpset() addsimps [Memrel_0, ordertype_0_sum_eq,  | 
| 1461 | 380  | 
ordertype_Memrel, well_ord_Memrel]) 1);  | 
| 849 | 381  | 
qed "oadd_0_left";  | 
382  | 
||
| 2469 | 383  | 
Addsimps [oadd_0, oadd_0_left];  | 
| 849 | 384  | 
|
385  | 
(*** Further properties of ordinal addition. Statements by Grabczewski,  | 
|
386  | 
proofs by lcp. ***)  | 
|
387  | 
||
| 5137 | 388  | 
Goalw [oadd_def] "[| k<i; Ord(j) |] ==> k < i++j";  | 
| 1461 | 389  | 
by (rtac ltE 1 THEN assume_tac 1);  | 
390  | 
by (rtac ltI 1);  | 
|
| 849 | 391  | 
by (REPEAT (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel] 2));  | 
392  | 
by (asm_simp_tac  | 
|
| 4091 | 393  | 
(simpset() addsimps [ordertype_pred_unfold,  | 
| 2493 | 394  | 
well_ord_radd, well_ord_Memrel,  | 
395  | 
ordertype_pred_Inl_eq,  | 
|
396  | 
lt_pred_Memrel, leI RS le_ordertype_Memrel]  | 
|
| 2469 | 397  | 
setloop rtac (InlI RSN (2,bexI))) 1);  | 
| 849 | 398  | 
qed "lt_oadd1";  | 
399  | 
||
| 
1032
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
400  | 
(*Thus also we obtain the rule i++j = k ==> i le k *)  | 
| 5137 | 401  | 
Goal "[| Ord(i); Ord(j) |] ==> i le i++j";  | 
| 1461 | 402  | 
by (rtac all_lt_imp_le 1);  | 
| 849 | 403  | 
by (REPEAT (ares_tac [Ord_oadd, lt_oadd1] 1));  | 
404  | 
qed "oadd_le_self";  | 
|
405  | 
||
406  | 
(** A couple of strange but necessary results! **)  | 
|
407  | 
||
| 5268 | 408  | 
Goal "A<=B ==> id(A) : ord_iso(A, Memrel(A), A, Memrel(B))";  | 
| 849 | 409  | 
by (resolve_tac [id_bij RS ord_isoI] 1);  | 
| 4091 | 410  | 
by (asm_simp_tac (simpset() addsimps [id_conv, Memrel_iff]) 1);  | 
| 2925 | 411  | 
by (Blast_tac 1);  | 
| 849 | 412  | 
qed "id_ord_iso_Memrel";  | 
413  | 
||
| 5268 | 414  | 
Goal "[| well_ord(A,r); k<j |] ==> \  | 
| 1461 | 415  | 
\ ordertype(A+k, radd(A, r, k, Memrel(j))) = \  | 
| 849 | 416  | 
\ ordertype(A+k, radd(A, r, k, Memrel(k)))";  | 
| 1461 | 417  | 
by (etac ltE 1);  | 
| 849 | 418  | 
by (resolve_tac [ord_iso_refl RS sum_ord_iso_cong RS ordertype_eq] 1);  | 
419  | 
by (eresolve_tac [OrdmemD RS id_ord_iso_Memrel RS ord_iso_sym] 1);  | 
|
420  | 
by (REPEAT_FIRST (ares_tac [well_ord_radd, well_ord_Memrel]));  | 
|
421  | 
qed "ordertype_sum_Memrel";  | 
|
422  | 
||
| 5137 | 423  | 
Goalw [oadd_def] "[| k<j; Ord(i) |] ==> i++k < i++j";  | 
| 1461 | 424  | 
by (rtac ltE 1 THEN assume_tac 1);  | 
| 849 | 425  | 
by (resolve_tac [ordertype_pred_unfold RS equalityD2 RS subsetD RS ltI] 1);  | 
426  | 
by (REPEAT_FIRST (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel]));  | 
|
| 1461 | 427  | 
by (rtac RepFun_eqI 1);  | 
428  | 
by (etac InrI 2);  | 
|
| 849 | 429  | 
by (asm_simp_tac  | 
| 4091 | 430  | 
(simpset() addsimps [ordertype_pred_Inr_eq, well_ord_Memrel,  | 
| 1461 | 431  | 
lt_pred_Memrel, leI RS le_ordertype_Memrel,  | 
432  | 
ordertype_sum_Memrel]) 1);  | 
|
| 849 | 433  | 
qed "oadd_lt_mono2";  | 
434  | 
||
| 5268 | 435  | 
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
 | 
436  | 
by (rtac Ord_linear_lt 1);  | 
| 
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
437  | 
by (REPEAT_SOME assume_tac);  | 
| 
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
438  | 
by (ALLGOALS  | 
| 4091 | 439  | 
(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
 | 
440  | 
qed "oadd_lt_cancel2";  | 
| 
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
441  | 
|
| 5268 | 442  | 
Goal "[| Ord(i); Ord(j); Ord(k) |] ==> i++j < i++k <-> j<k";  | 
| 4091 | 443  | 
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
 | 
444  | 
qed "oadd_lt_iff2";  | 
| 
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
445  | 
|
| 5268 | 446  | 
Goal "[| i++j = i++k; Ord(i); Ord(j); Ord(k) |] ==> j=k";  | 
| 849 | 447  | 
by (rtac Ord_linear_lt 1);  | 
448  | 
by (REPEAT_SOME assume_tac);  | 
|
449  | 
by (ALLGOALS  | 
|
| 4091 | 450  | 
(fast_tac (claset() addDs [oadd_lt_mono2]  | 
451  | 
addss (simpset() addsimps [lt_not_refl]))));  | 
|
| 849 | 452  | 
qed "oadd_inject";  | 
453  | 
||
| 5067 | 454  | 
Goalw [oadd_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
455  | 
"[| k < i++j; Ord(i); Ord(j) |] ==> k<i | (EX l:j. k = i++l )";  | 
| 849 | 456  | 
(*Rotate the hypotheses so that simplification will work*)  | 
457  | 
by (etac revcut_rl 1);  | 
|
458  | 
by (asm_full_simp_tac  | 
|
| 4091 | 459  | 
(simpset() addsimps [ordertype_pred_unfold, well_ord_radd,  | 
| 1461 | 460  | 
well_ord_Memrel]) 1);  | 
| 849 | 461  | 
by (eresolve_tac [ltD RS RepFunE] 1);  | 
| 4091 | 462  | 
by (fast_tac (claset() addss  | 
463  | 
(simpset() addsimps [ordertype_pred_Inl_eq, well_ord_Memrel,  | 
|
| 1461 | 464  | 
ltI, lt_pred_Memrel, le_ordertype_Memrel, leI,  | 
465  | 
ordertype_pred_Inr_eq,  | 
|
466  | 
ordertype_sum_Memrel])) 1);  | 
|
| 849 | 467  | 
qed "lt_oadd_disj";  | 
468  | 
||
469  | 
||
470  | 
(*** Ordinal addition with successor -- via associativity! ***)  | 
|
471  | 
||
| 5067 | 472  | 
Goalw [oadd_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
473  | 
"[| Ord(i); Ord(j); Ord(k) |] ==> (i++j)++k = i++(j++k)";  | 
| 849 | 474  | 
by (resolve_tac [ordertype_eq RS trans] 1);  | 
475  | 
by (rtac ([ordertype_ord_iso RS ord_iso_sym, ord_iso_refl] MRS  | 
|
| 1461 | 476  | 
sum_ord_iso_cong) 1);  | 
| 849 | 477  | 
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel, Ord_ordertype] 1));  | 
478  | 
by (resolve_tac [sum_assoc_ord_iso RS ordertype_eq RS trans] 1);  | 
|
479  | 
by (rtac ([ord_iso_refl, ordertype_ord_iso] MRS sum_ord_iso_cong RS  | 
|
| 1461 | 480  | 
ordertype_eq) 2);  | 
| 849 | 481  | 
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel, Ord_ordertype] 1));  | 
482  | 
qed "oadd_assoc";  | 
|
483  | 
||
| 5268 | 484  | 
Goal "[| Ord(i);  Ord(j) |] ==> i++j = i Un (UN k:j. {i++k})";
 | 
| 849 | 485  | 
by (rtac (subsetI RS equalityI) 1);  | 
486  | 
by (eresolve_tac [ltI RS lt_oadd_disj RS disjE] 1);  | 
|
487  | 
by (REPEAT (ares_tac [Ord_oadd] 1));  | 
|
| 4091 | 488  | 
by (fast_tac (claset() addIs [lt_oadd1, oadd_lt_mono2]  | 
489  | 
addss (simpset() addsimps [Ord_mem_iff_lt, Ord_oadd])) 3);  | 
|
| 2925 | 490  | 
by (Blast_tac 2);  | 
| 4091 | 491  | 
by (blast_tac (claset() addSEs [ltE]) 1);  | 
| 849 | 492  | 
qed "oadd_unfold";  | 
493  | 
||
| 5137 | 494  | 
Goal "Ord(i) ==> i++1 = succ(i)";  | 
| 4091 | 495  | 
by (asm_simp_tac (simpset() addsimps [oadd_unfold, Ord_1, oadd_0]) 1);  | 
| 2925 | 496  | 
by (Blast_tac 1);  | 
| 849 | 497  | 
qed "oadd_1";  | 
498  | 
||
| 5268 | 499  | 
Goal "[| Ord(i); Ord(j) |] ==> i++succ(j) = succ(i++j)";  | 
| 6153 | 500  | 
(*FOL_ss prevents looping*)  | 
501  | 
by (asm_simp_tac (FOL_ss addsimps [Ord_oadd, oadd_1 RS sym]) 1);  | 
|
| 4091 | 502  | 
by (asm_simp_tac (simpset() addsimps [oadd_1, oadd_assoc, Ord_1]) 1);  | 
| 849 | 503  | 
qed "oadd_succ";  | 
504  | 
||
505  | 
||
506  | 
(** Ordinal addition with limit ordinals **)  | 
|
507  | 
||
| 5268 | 508  | 
val prems =  | 
509  | 
Goal "[| Ord(i); !!x. x:A ==> Ord(j(x)); a:A |] ==> \  | 
|
510  | 
\ i ++ (UN x:A. j(x)) = (UN x:A. i++j(x))";  | 
|
| 5618 | 511  | 
by (blast_tac (claset() addIs prems @ [ltI, Ord_UN, Ord_oadd,  | 
512  | 
lt_oadd1 RS ltD, oadd_lt_mono2 RS ltD]  | 
|
| 9173 | 513  | 
addSEs [ltE] addSDs [ltI RS lt_oadd_disj]) 1);  | 
| 849 | 514  | 
qed "oadd_UN";  | 
515  | 
||
| 5268 | 516  | 
Goal "[| Ord(i); Limit(j) |] ==> i++j = (UN k:j. i++k)";  | 
| 849 | 517  | 
by (forward_tac [Limit_has_0 RS ltD] 1);  | 
| 4091 | 518  | 
by (asm_simp_tac (simpset() addsimps [Limit_is_Ord RS Ord_in_Ord,  | 
| 9173 | 519  | 
oadd_UN RS sym, Union_eq_UN RS sym,  | 
520  | 
Limit_Union_eq]) 1);  | 
|
| 849 | 521  | 
qed "oadd_Limit";  | 
522  | 
||
523  | 
(** Order/monotonicity properties of ordinal addition **)  | 
|
524  | 
||
| 5137 | 525  | 
Goal "[| Ord(i); Ord(j) |] ==> i le j++i";  | 
| 849 | 526  | 
by (eres_inst_tac [("i","i")] trans_induct3 1);
 | 
| 4091 | 527  | 
by (asm_simp_tac (simpset() addsimps [Ord_0_le]) 1);  | 
528  | 
by (asm_simp_tac (simpset() addsimps [oadd_succ, succ_leI]) 1);  | 
|
529  | 
by (asm_simp_tac (simpset() addsimps [oadd_Limit]) 1);  | 
|
| 1461 | 530  | 
by (rtac le_trans 1);  | 
531  | 
by (rtac le_implies_UN_le_UN 2);  | 
|
| 2925 | 532  | 
by (Blast_tac 2);  | 
| 4091 | 533  | 
by (asm_simp_tac (simpset() addsimps [Union_eq_UN RS sym, Limit_Union_eq,  | 
| 2493 | 534  | 
le_refl, Limit_is_Ord]) 1);  | 
| 849 | 535  | 
qed "oadd_le_self2";  | 
536  | 
||
| 5137 | 537  | 
Goal "[| k le j; Ord(i) |] ==> k++i le j++i";  | 
| 7499 | 538  | 
by (ftac lt_Ord 1);  | 
539  | 
by (ftac le_Ord2 1);  | 
|
| 1461 | 540  | 
by (etac trans_induct3 1);  | 
| 2469 | 541  | 
by (Asm_simp_tac 1);  | 
| 4091 | 542  | 
by (asm_simp_tac (simpset() addsimps [oadd_succ, succ_le_iff]) 1);  | 
543  | 
by (asm_simp_tac (simpset() addsimps [oadd_Limit]) 1);  | 
|
| 1461 | 544  | 
by (rtac le_implies_UN_le_UN 1);  | 
| 2925 | 545  | 
by (Blast_tac 1);  | 
| 849 | 546  | 
qed "oadd_le_mono1";  | 
547  | 
||
| 5137 | 548  | 
Goal "[| i' le i; j'<j |] ==> i'++j' < i++j";  | 
| 1461 | 549  | 
by (rtac lt_trans1 1);  | 
| 849 | 550  | 
by (REPEAT (eresolve_tac [asm_rl, oadd_le_mono1, oadd_lt_mono2, ltE,  | 
| 1461 | 551  | 
Ord_succD] 1));  | 
| 849 | 552  | 
qed "oadd_lt_mono";  | 
553  | 
||
| 5137 | 554  | 
Goal "[| i' le i; j' le j |] ==> i'++j' le i++j";  | 
| 4091 | 555  | 
by (asm_simp_tac (simpset() addsimps [oadd_succ RS sym, le_Ord2, oadd_lt_mono]) 1);  | 
| 849 | 556  | 
qed "oadd_le_mono";  | 
557  | 
||
| 5268 | 558  | 
Goal "[| Ord(i); Ord(j); Ord(k) |] ==> i++j le i++k <-> j le k";  | 
| 4091 | 559  | 
by (asm_simp_tac (simpset() addsimps [oadd_lt_iff2, oadd_succ RS sym,  | 
| 1461 | 560  | 
Ord_succ]) 1);  | 
| 
1032
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
561  | 
qed "oadd_le_iff2";  | 
| 
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
562  | 
|
| 849 | 563  | 
|
564  | 
(** Ordinal subtraction; the difference is ordertype(j-i, Memrel(j)).  | 
|
565  | 
Probably simpler to define the difference recursively!  | 
|
566  | 
**)  | 
|
567  | 
||
| 5268 | 568  | 
Goal "A<=B ==> (lam y:B. if(y:A, Inl(y), Inr(y))) : bij(B, A+(B-A))";  | 
| 3840 | 569  | 
by (res_inst_tac [("d", "case(%x. x, %y. y)")] lam_bijective 1);
 | 
| 4091 | 570  | 
by (blast_tac (claset() addSIs [if_type]) 1);  | 
571  | 
by (fast_tac (claset() addSIs [case_type]) 1);  | 
|
| 1461 | 572  | 
by (etac sumE 2);  | 
| 5137 | 573  | 
by (ALLGOALS Asm_simp_tac);  | 
| 849 | 574  | 
qed "bij_sum_Diff";  | 
575  | 
||
| 5268 | 576  | 
Goal "i le j ==> \  | 
| 1461 | 577  | 
\ ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j))) = \  | 
| 849 | 578  | 
\ ordertype(j, Memrel(j))";  | 
| 4091 | 579  | 
by (safe_tac (claset() addSDs [le_subset_iff RS iffD1]));  | 
| 849 | 580  | 
by (resolve_tac [bij_sum_Diff RS ord_isoI RS ord_iso_sym RS ordertype_eq] 1);  | 
| 1461 | 581  | 
by (etac well_ord_Memrel 3);  | 
| 849 | 582  | 
by (assume_tac 1);  | 
| 5137 | 583  | 
by (asm_simp_tac (simpset() addsimps [Memrel_iff]) 1);  | 
| 849 | 584  | 
by (forw_inst_tac [("j", "y")] Ord_in_Ord 1 THEN assume_tac 1);
 | 
585  | 
by (forw_inst_tac [("j", "x")] Ord_in_Ord 1 THEN assume_tac 1);
 | 
|
| 4091 | 586  | 
by (asm_simp_tac (simpset() addsimps [Ord_mem_iff_lt, lt_Ord, not_lt_iff_le]) 1);  | 
587  | 
by (blast_tac (claset() addIs [lt_trans2, lt_trans]) 1);  | 
|
| 849 | 588  | 
qed "ordertype_sum_Diff";  | 
589  | 
||
| 5067 | 590  | 
Goalw [oadd_def, odiff_def]  | 
| 5268 | 591  | 
"i le j \  | 
592  | 
\ ==> i ++ (j--i) = ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j)))";  | 
|
| 4091 | 593  | 
by (safe_tac (claset() addSDs [le_subset_iff RS iffD1]));  | 
| 849 | 594  | 
by (resolve_tac [sum_ord_iso_cong RS ordertype_eq] 1);  | 
| 1461 | 595  | 
by (etac id_ord_iso_Memrel 1);  | 
| 849 | 596  | 
by (resolve_tac [ordertype_ord_iso RS ord_iso_sym] 1);  | 
597  | 
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel RS well_ord_subset,  | 
|
| 1461 | 598  | 
Diff_subset] 1));  | 
| 849 | 599  | 
qed "oadd_ordertype_Diff";  | 
600  | 
||
| 5137 | 601  | 
Goal "i le j ==> i ++ (j--i) = j";  | 
| 4091 | 602  | 
by (asm_simp_tac (simpset() addsimps [oadd_ordertype_Diff, ordertype_sum_Diff,  | 
| 1461 | 603  | 
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
 | 
604  | 
qed "oadd_odiff_inverse";  | 
| 
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
605  | 
|
| 5067 | 606  | 
Goalw [odiff_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
607  | 
"[| Ord(i); Ord(j) |] ==> Ord(i--j)";  | 
| 
1032
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
608  | 
by (REPEAT (ares_tac [Ord_ordertype, well_ord_Memrel RS well_ord_subset,  | 
| 1461 | 609  | 
Diff_subset] 1));  | 
| 
1032
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
610  | 
qed "Ord_odiff";  | 
| 849 | 611  | 
|
| 
1032
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
612  | 
(*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
 | 
613  | 
i++j = k ==> j = k--i. *)  | 
| 5268 | 614  | 
Goal "[| Ord(i); Ord(j) |] ==> (i++j) -- i = j";  | 
| 1461 | 615  | 
by (rtac oadd_inject 1);  | 
| 
1032
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
616  | 
by (REPEAT (ares_tac [Ord_ordertype, Ord_oadd, Ord_odiff] 2));  | 
| 4091 | 617  | 
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
 | 
618  | 
qed "odiff_oadd_inverse";  | 
| 
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
619  | 
|
| 
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
620  | 
val [i_lt_j, k_le_i] = goal OrderType.thy  | 
| 
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
621  | 
"[| 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
 | 
622  | 
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
 | 
623  | 
by (simp_tac  | 
| 4091 | 624  | 
(simpset() addsimps [i_lt_j, k_le_i, [k_le_i, leI] MRS le_trans,  | 
| 5268 | 625  | 
oadd_odiff_inverse]) 1);  | 
| 
1032
 
54b9f670c67e
Proved odiff_oadd_inverse, oadd_lt_cancel2, oadd_lt_iff2,
 
lcp 
parents: 
984 
diff
changeset
 | 
626  | 
by (REPEAT (resolve_tac (Ord_odiff ::  | 
| 1461 | 627  | 
([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
 | 
628  | 
qed "odiff_lt_mono2";  | 
| 849 | 629  | 
|
630  | 
||
631  | 
(**** Ordinal Multiplication ****)  | 
|
632  | 
||
| 5067 | 633  | 
Goalw [omult_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
634  | 
"[| Ord(i); Ord(j) |] ==> Ord(i**j)";  | 
| 849 | 635  | 
by (REPEAT (ares_tac [Ord_ordertype, well_ord_rmult, well_ord_Memrel] 1));  | 
636  | 
qed "Ord_omult";  | 
|
637  | 
||
638  | 
(*** A useful unfolding law ***)  | 
|
639  | 
||
| 5067 | 640  | 
Goalw [pred_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
641  | 
"[| a:A; b:B |] ==> \  | 
| 1461 | 642  | 
\ pred(A*B, <a,b>, rmult(A,r,B,s)) = \  | 
| 849 | 643  | 
\        pred(A,a,r)*B Un ({a} * pred(B,b,s))";
 | 
| 2493 | 644  | 
by (rtac equalityI 1);  | 
| 4152 | 645  | 
by Safe_tac;  | 
| 4091 | 646  | 
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [rmult_iff])));  | 
| 2925 | 647  | 
by (ALLGOALS (Blast_tac));  | 
| 849 | 648  | 
qed "pred_Pair_eq";  | 
649  | 
||
| 5268 | 650  | 
Goal "[| a:A; b:B; well_ord(A,r); well_ord(B,s) |] ==> \  | 
| 849 | 651  | 
\ ordertype(pred(A*B, <a,b>, rmult(A,r,B,s)), rmult(A,r,B,s)) = \  | 
| 3016 | 652  | 
\ ordertype(pred(A,a,r)*B + pred(B,b,s), \  | 
| 849 | 653  | 
\ radd(A*B, rmult(A,r,B,s), B, s))";  | 
| 4091 | 654  | 
by (asm_simp_tac (simpset() addsimps [pred_Pair_eq]) 1);  | 
| 849 | 655  | 
by (resolve_tac [ordertype_eq RS sym] 1);  | 
| 1461 | 656  | 
by (rtac prod_sum_singleton_ord_iso 1);  | 
| 984 | 657  | 
by (REPEAT_FIRST (ares_tac [pred_subset, well_ord_rmult RS well_ord_subset]));  | 
| 4091 | 658  | 
by (blast_tac (claset() addSEs [predE]) 1);  | 
| 849 | 659  | 
qed "ordertype_pred_Pair_eq";  | 
660  | 
||
| 5067 | 661  | 
Goalw [oadd_def, omult_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
662  | 
"[| i'<i; j'<j |] ==> \  | 
| 849 | 663  | 
\ ordertype(pred(i*j, <i',j'>, rmult(i,Memrel(i),j,Memrel(j))), \  | 
| 3016 | 664  | 
\ rmult(i,Memrel(i),j,Memrel(j))) = \  | 
| 849 | 665  | 
\ j**i' ++ j'";  | 
| 4091 | 666  | 
by (asm_simp_tac (simpset() addsimps [ordertype_pred_Pair_eq, lt_pred_Memrel,  | 
| 3016 | 667  | 
ltD, lt_Ord2, well_ord_Memrel]) 1);  | 
| 1461 | 668  | 
by (rtac trans 1);  | 
| 849 | 669  | 
by (resolve_tac [ordertype_ord_iso RS sum_ord_iso_cong RS ordertype_eq] 2);  | 
| 1461 | 670  | 
by (rtac ord_iso_refl 3);  | 
| 849 | 671  | 
by (resolve_tac [id_bij RS ord_isoI RS ordertype_eq] 1);  | 
672  | 
by (REPEAT_FIRST (eresolve_tac [SigmaE, sumE, ltE, ssubst]));  | 
|
673  | 
by (REPEAT_FIRST (ares_tac [well_ord_rmult, well_ord_radd, well_ord_Memrel,  | 
|
| 1461 | 674  | 
Ord_ordertype]));  | 
| 4091 | 675  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [Memrel_iff])));  | 
| 4152 | 676  | 
by Safe_tac;  | 
| 4091 | 677  | 
by (ALLGOALS (blast_tac (claset() addIs [Ord_trans])));  | 
| 849 | 678  | 
qed "ordertype_pred_Pair_lemma";  | 
679  | 
||
| 5067 | 680  | 
Goalw [omult_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
681  | 
"[| Ord(i); Ord(j); k<j**i |] ==> \  | 
| 849 | 682  | 
\ EX j' i'. k = j**i' ++ j' & j'<j & i'<i";  | 
| 4091 | 683  | 
by (asm_full_simp_tac (simpset() addsimps [ordertype_pred_unfold,  | 
| 1461 | 684  | 
well_ord_rmult, well_ord_Memrel]) 1);  | 
| 4091 | 685  | 
by (safe_tac (claset() addSEs [ltE]));  | 
686  | 
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
 | 
687  | 
symmetric omult_def]) 1);  | 
| 4091 | 688  | 
by (blast_tac (claset() addIs [ltI]) 1);  | 
| 849 | 689  | 
qed "lt_omult";  | 
690  | 
||
| 5067 | 691  | 
Goalw [omult_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
692  | 
"[| j'<j; i'<i |] ==> j**i' ++ j' < j**i";  | 
| 1461 | 693  | 
by (rtac ltI 1);  | 
| 984 | 694  | 
by (asm_simp_tac  | 
| 4091 | 695  | 
(simpset() addsimps [Ord_ordertype, well_ord_rmult, well_ord_Memrel,  | 
| 2493 | 696  | 
lt_Ord2]) 2);  | 
| 984 | 697  | 
by (asm_simp_tac  | 
| 4091 | 698  | 
(simpset() addsimps [ordertype_pred_unfold,  | 
| 1461 | 699  | 
well_ord_rmult, well_ord_Memrel, lt_Ord2]) 1);  | 
| 2469 | 700  | 
by (rtac bexI 1);  | 
| 4091 | 701  | 
by (blast_tac (claset() addSEs [ltE]) 2);  | 
| 849 | 702  | 
by (asm_simp_tac  | 
| 4091 | 703  | 
(simpset() addsimps [ordertype_pred_Pair_lemma, ltI,  | 
| 2493 | 704  | 
symmetric omult_def]) 1);  | 
| 849 | 705  | 
qed "omult_oadd_lt";  | 
706  | 
||
| 5268 | 707  | 
Goal "[| Ord(i);  Ord(j) |] ==> j**i = (UN j':j. UN i':i. {j**i' ++ j'})";
 | 
| 849 | 708  | 
by (rtac (subsetI RS equalityI) 1);  | 
709  | 
by (resolve_tac [lt_omult RS exE] 1);  | 
|
| 1461 | 710  | 
by (etac ltI 3);  | 
| 849 | 711  | 
by (REPEAT (ares_tac [Ord_omult] 1));  | 
| 4091 | 712  | 
by (blast_tac (claset() addSEs [ltE]) 1);  | 
713  | 
by (blast_tac (claset() addIs [omult_oadd_lt RS ltD, ltI]) 1);  | 
|
| 849 | 714  | 
qed "omult_unfold";  | 
715  | 
||
716  | 
(*** Basic laws for ordinal multiplication ***)  | 
|
717  | 
||
718  | 
(** Ordinal multiplication by zero **)  | 
|
719  | 
||
| 5067 | 720  | 
Goalw [omult_def] "i**0 = 0";  | 
| 2469 | 721  | 
by (Asm_simp_tac 1);  | 
| 849 | 722  | 
qed "omult_0";  | 
723  | 
||
| 5067 | 724  | 
Goalw [omult_def] "0**i = 0";  | 
| 2469 | 725  | 
by (Asm_simp_tac 1);  | 
| 849 | 726  | 
qed "omult_0_left";  | 
727  | 
||
| 2469 | 728  | 
Addsimps [omult_0, omult_0_left];  | 
729  | 
||
| 849 | 730  | 
(** Ordinal multiplication by 1 **)  | 
731  | 
||
| 5137 | 732  | 
Goalw [omult_def] "Ord(i) ==> i**1 = i";  | 
| 849 | 733  | 
by (resolve_tac [ord_isoI RS ordertype_eq RS trans] 1);  | 
734  | 
by (res_inst_tac [("c", "snd"), ("d", "%z.<0,z>")] lam_bijective 1);
 | 
|
735  | 
by (REPEAT_FIRST (eresolve_tac [snd_type, SigmaE, succE, emptyE,  | 
|
| 1461 | 736  | 
well_ord_Memrel, ordertype_Memrel]));  | 
| 4091 | 737  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [rmult_iff, Memrel_iff])));  | 
| 849 | 738  | 
qed "omult_1";  | 
739  | 
||
| 5137 | 740  | 
Goalw [omult_def] "Ord(i) ==> 1**i = i";  | 
| 849 | 741  | 
by (resolve_tac [ord_isoI RS ordertype_eq RS trans] 1);  | 
742  | 
by (res_inst_tac [("c", "fst"), ("d", "%z.<z,0>")] lam_bijective 1);
 | 
|
743  | 
by (REPEAT_FIRST (eresolve_tac [fst_type, SigmaE, succE, emptyE,  | 
|
| 1461 | 744  | 
well_ord_Memrel, ordertype_Memrel]));  | 
| 4091 | 745  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [rmult_iff, Memrel_iff])));  | 
| 849 | 746  | 
qed "omult_1_left";  | 
747  | 
||
| 2469 | 748  | 
Addsimps [omult_1, omult_1_left];  | 
749  | 
||
| 849 | 750  | 
(** Distributive law for ordinal multiplication and addition **)  | 
751  | 
||
| 5067 | 752  | 
Goalw [omult_def, oadd_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
753  | 
"[| Ord(i); Ord(j); Ord(k) |] ==> i**(j++k) = (i**j)++(i**k)";  | 
| 849 | 754  | 
by (resolve_tac [ordertype_eq RS trans] 1);  | 
755  | 
by (rtac ([ordertype_ord_iso RS ord_iso_sym, ord_iso_refl] MRS  | 
|
| 1461 | 756  | 
prod_ord_iso_cong) 1);  | 
| 849 | 757  | 
by (REPEAT (ares_tac [well_ord_rmult, well_ord_radd, well_ord_Memrel,  | 
| 1461 | 758  | 
Ord_ordertype] 1));  | 
| 849 | 759  | 
by (rtac (sum_prod_distrib_ord_iso RS ordertype_eq RS trans) 1);  | 
760  | 
by (rtac ordertype_eq 2);  | 
|
761  | 
by (rtac ([ordertype_ord_iso, ordertype_ord_iso] MRS sum_ord_iso_cong) 2);  | 
|
762  | 
by (REPEAT (ares_tac [well_ord_rmult, well_ord_radd, well_ord_Memrel,  | 
|
| 1461 | 763  | 
Ord_ordertype] 1));  | 
| 849 | 764  | 
qed "oadd_omult_distrib";  | 
765  | 
||
| 5137 | 766  | 
Goal "[| Ord(i); Ord(j) |] ==> i**succ(j) = (i**j)++i";  | 
| 6153 | 767  | 
(*FOL_ss prevents looping*)  | 
768  | 
by (asm_simp_tac (FOL_ss addsimps [oadd_1 RS sym]) 1);  | 
|
| 849 | 769  | 
by (asm_simp_tac  | 
| 4091 | 770  | 
(simpset() addsimps [omult_1, oadd_omult_distrib, Ord_1]) 1);  | 
| 849 | 771  | 
qed "omult_succ";  | 
772  | 
||
773  | 
(** Associative law **)  | 
|
774  | 
||
| 5067 | 775  | 
Goalw [omult_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
776  | 
"[| Ord(i); Ord(j); Ord(k) |] ==> (i**j)**k = i**(j**k)";  | 
| 849 | 777  | 
by (resolve_tac [ordertype_eq RS trans] 1);  | 
778  | 
by (rtac ([ord_iso_refl, ordertype_ord_iso RS ord_iso_sym] MRS  | 
|
| 1461 | 779  | 
prod_ord_iso_cong) 1);  | 
| 849 | 780  | 
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1));  | 
781  | 
by (resolve_tac [prod_assoc_ord_iso RS ord_iso_sym RS  | 
|
| 1461 | 782  | 
ordertype_eq RS trans] 1);  | 
| 849 | 783  | 
by (rtac ([ordertype_ord_iso, ord_iso_refl] MRS prod_ord_iso_cong RS  | 
| 1461 | 784  | 
ordertype_eq) 2);  | 
| 849 | 785  | 
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel, Ord_ordertype] 1));  | 
786  | 
qed "omult_assoc";  | 
|
787  | 
||
788  | 
||
789  | 
(** Ordinal multiplication with limit ordinals **)  | 
|
790  | 
||
| 5268 | 791  | 
val prems =  | 
792  | 
Goal "[| Ord(i); !!x. x:A ==> Ord(j(x)) |] ==> \  | 
|
793  | 
\ i ** (UN x:A. j(x)) = (UN x:A. i**j(x))";  | 
|
| 5529 | 794  | 
by (asm_simp_tac (simpset() addsimps prems @ [Ord_UN, omult_unfold]) 1);  | 
| 2925 | 795  | 
by (Blast_tac 1);  | 
| 849 | 796  | 
qed "omult_UN";  | 
| 467 | 797  | 
|
| 5268 | 798  | 
Goal "[| Ord(i); Limit(j) |] ==> i**j = (UN k:j. i**k)";  | 
| 849 | 799  | 
by (asm_simp_tac  | 
| 4091 | 800  | 
(simpset() addsimps [Limit_is_Ord RS Ord_in_Ord, omult_UN RS sym,  | 
| 1461 | 801  | 
Union_eq_UN RS sym, Limit_Union_eq]) 1);  | 
| 849 | 802  | 
qed "omult_Limit";  | 
803  | 
||
804  | 
||
805  | 
(*** Ordering/monotonicity properties of ordinal multiplication ***)  | 
|
806  | 
||
807  | 
(*As a special case we have "[| 0<i; 0<j |] ==> 0 < i**j" *)  | 
|
| 5137 | 808  | 
Goal "[| k<i; 0<j |] ==> k < i**j";  | 
| 4091 | 809  | 
by (safe_tac (claset() addSEs [ltE] addSIs [ltI, Ord_omult]));  | 
810  | 
by (asm_simp_tac (simpset() addsimps [omult_unfold]) 1);  | 
|
| 2469 | 811  | 
by (REPEAT_FIRST (ares_tac [bexI]));  | 
812  | 
by (Asm_simp_tac 1);  | 
|
| 849 | 813  | 
qed "lt_omult1";  | 
814  | 
||
| 5137 | 815  | 
Goal "[| Ord(i); 0<j |] ==> i le i**j";  | 
| 1461 | 816  | 
by (rtac all_lt_imp_le 1);  | 
| 849 | 817  | 
by (REPEAT (ares_tac [Ord_omult, lt_omult1, lt_Ord2] 1));  | 
818  | 
qed "omult_le_self";  | 
|
819  | 
||
| 5137 | 820  | 
Goal "[| k le j; Ord(i) |] ==> k**i le j**i";  | 
| 7499 | 821  | 
by (ftac lt_Ord 1);  | 
822  | 
by (ftac le_Ord2 1);  | 
|
| 1461 | 823  | 
by (etac trans_induct3 1);  | 
| 4091 | 824  | 
by (asm_simp_tac (simpset() addsimps [le_refl, Ord_0]) 1);  | 
825  | 
by (asm_simp_tac (simpset() addsimps [omult_succ, oadd_le_mono]) 1);  | 
|
826  | 
by (asm_simp_tac (simpset() addsimps [omult_Limit]) 1);  | 
|
| 1461 | 827  | 
by (rtac le_implies_UN_le_UN 1);  | 
| 2925 | 828  | 
by (Blast_tac 1);  | 
| 849 | 829  | 
qed "omult_le_mono1";  | 
830  | 
||
| 5137 | 831  | 
Goal "[| k<j; 0<i |] ==> i**k < i**j";  | 
| 1461 | 832  | 
by (rtac ltI 1);  | 
| 4091 | 833  | 
by (asm_simp_tac (simpset() addsimps [omult_unfold, lt_Ord2]) 1);  | 
834  | 
by (safe_tac (claset() addSEs [ltE] addSIs [Ord_omult]));  | 
|
| 2469 | 835  | 
by (REPEAT_FIRST (ares_tac [bexI]));  | 
| 4091 | 836  | 
by (asm_simp_tac (simpset() addsimps [Ord_omult]) 1);  | 
| 849 | 837  | 
qed "omult_lt_mono2";  | 
838  | 
||
| 5137 | 839  | 
Goal "[| k le j; Ord(i) |] ==> i**k le i**j";  | 
| 1461 | 840  | 
by (rtac subset_imp_le 1);  | 
| 4091 | 841  | 
by (safe_tac (claset() addSEs [ltE, make_elim Ord_succD] addSIs [Ord_omult]));  | 
842  | 
by (asm_full_simp_tac (simpset() addsimps [omult_unfold]) 1);  | 
|
843  | 
by (deepen_tac (claset() addEs [Ord_trans]) 0 1);  | 
|
| 849 | 844  | 
qed "omult_le_mono2";  | 
845  | 
||
| 5137 | 846  | 
Goal "[| i' le i; j' le j |] ==> i'**j' le i**j";  | 
| 1461 | 847  | 
by (rtac le_trans 1);  | 
| 849 | 848  | 
by (REPEAT (eresolve_tac [asm_rl, omult_le_mono1, omult_le_mono2, ltE,  | 
| 1461 | 849  | 
Ord_succD] 1));  | 
| 849 | 850  | 
qed "omult_le_mono";  | 
851  | 
||
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
852  | 
Goal "[| i' le i; j'<j; 0<i |] ==> i'**j' < i**j";  | 
| 1461 | 853  | 
by (rtac lt_trans1 1);  | 
| 849 | 854  | 
by (REPEAT (eresolve_tac [asm_rl, omult_le_mono1, omult_lt_mono2, ltE,  | 
| 1461 | 855  | 
Ord_succD] 1));  | 
| 849 | 856  | 
qed "omult_lt_mono";  | 
857  | 
||
| 5137 | 858  | 
Goal "[| Ord(i); 0<j |] ==> i le j**i";  | 
| 7499 | 859  | 
by (ftac lt_Ord2 1);  | 
| 849 | 860  | 
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
 | 
861  | 
by (Asm_simp_tac 1);  | 
| 
 
68c6159440f1
new lemmas for Ntree recursor example;  more simprules;  more lemmas borrowed
 
paulson 
parents: 
7499 
diff
changeset
 | 
862  | 
by (asm_simp_tac (simpset() addsimps [omult_succ]) 1);  | 
| 1461 | 863  | 
by (etac lt_trans1 1);  | 
| 849 | 864  | 
by (res_inst_tac [("b", "j**x")] (oadd_0 RS subst) 1 THEN 
 | 
865  | 
rtac oadd_lt_mono2 2);  | 
|
866  | 
by (REPEAT (ares_tac [Ord_omult] 1));  | 
|
| 4091 | 867  | 
by (asm_simp_tac (simpset() addsimps [omult_Limit]) 1);  | 
| 1461 | 868  | 
by (rtac le_trans 1);  | 
869  | 
by (rtac le_implies_UN_le_UN 2);  | 
|
| 2925 | 870  | 
by (Blast_tac 2);  | 
| 4091 | 871  | 
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
 | 
872  | 
Limit_is_Ord]) 1);  | 
| 849 | 873  | 
qed "omult_le_self2";  | 
874  | 
||
875  | 
||
876  | 
(** Further properties of ordinal multiplication **)  | 
|
877  | 
||
| 5137 | 878  | 
Goal "[| i**j = i**k; 0<i; Ord(j); Ord(k) |] ==> j=k";  | 
| 849 | 879  | 
by (rtac Ord_linear_lt 1);  | 
880  | 
by (REPEAT_SOME assume_tac);  | 
|
881  | 
by (ALLGOALS  | 
|
| 4091 | 882  | 
(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
 | 
883  | 
addss (simpset() addsimps [lt_not_refl]))));  | 
| 849 | 884  | 
qed "omult_inject";  | 
885  | 
||
886  |