| author | wenzelm | 
| Tue, 18 May 1999 15:52:34 +0200 | |
| changeset 6671 | 677713791bd8 | 
| parent 6068 | 2d8f3e1f1151 | 
| child 8201 | a81d18b0a9b1 | 
| permissions | -rw-r--r-- | 
| 1461 | 1  | 
(* Title: ZF/OrderArith.ML  | 
| 437 | 2  | 
ID: $Id$  | 
| 1461 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 437 | 4  | 
Copyright 1994 University of Cambridge  | 
5  | 
||
6  | 
Towards ordinal arithmetic  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
open OrderArith;  | 
|
10  | 
||
11  | 
||
12  | 
(**** Addition of relations -- disjoint sum ****)  | 
|
13  | 
||
14  | 
(** Rewrite rules. Can be used to obtain introduction rules **)  | 
|
15  | 
||
| 5067 | 16  | 
Goalw [radd_def]  | 
| 437 | 17  | 
"<Inl(a), Inr(b)> : radd(A,r,B,s) <-> a:A & b:B";  | 
| 2925 | 18  | 
by (Blast_tac 1);  | 
| 760 | 19  | 
qed "radd_Inl_Inr_iff";  | 
| 437 | 20  | 
|
| 5067 | 21  | 
Goalw [radd_def]  | 
| 859 | 22  | 
"<Inl(a'), Inl(a)> : radd(A,r,B,s) <-> a':A & a:A & <a',a>:r";  | 
| 2925 | 23  | 
by (Blast_tac 1);  | 
| 760 | 24  | 
qed "radd_Inl_iff";  | 
| 437 | 25  | 
|
| 5067 | 26  | 
Goalw [radd_def]  | 
| 859 | 27  | 
"<Inr(b'), Inr(b)> : radd(A,r,B,s) <-> b':B & b:B & <b',b>:s";  | 
| 2925 | 28  | 
by (Blast_tac 1);  | 
| 760 | 29  | 
qed "radd_Inr_iff";  | 
| 437 | 30  | 
|
| 5067 | 31  | 
Goalw [radd_def]  | 
| 437 | 32  | 
"<Inr(b), Inl(a)> : radd(A,r,B,s) <-> False";  | 
| 2925 | 33  | 
by (Blast_tac 1);  | 
| 760 | 34  | 
qed "radd_Inr_Inl_iff";  | 
| 437 | 35  | 
|
36  | 
(** Elimination Rule **)  | 
|
37  | 
||
| 5268 | 38  | 
val major::prems = Goalw [radd_def]  | 
| 1461 | 39  | 
"[| <p',p> : radd(A,r,B,s); \  | 
40  | 
\ !!x y. [| p'=Inl(x); x:A; p=Inr(y); y:B |] ==> Q; \  | 
|
41  | 
\ !!x' x. [| p'=Inl(x'); p=Inl(x); <x',x>: r; x':A; x:A |] ==> Q; \  | 
|
42  | 
\ !!y' y. [| p'=Inr(y'); p=Inr(y); <y',y>: s; y':B; y:B |] ==> Q \  | 
|
| 437 | 43  | 
\ |] ==> Q";  | 
44  | 
by (cut_facts_tac [major] 1);  | 
|
45  | 
(*Split into the three cases*)  | 
|
| 2925 | 46  | 
by (REPEAT_FIRST (*can't use safe_tac: don't want hyp_subst_tac*)  | 
| 437 | 47  | 
(eresolve_tac [CollectE, Pair_inject, conjE, exE, SigmaE, disjE]));  | 
48  | 
(*Apply each premise to correct subgoal; can't just use fast_tac  | 
|
49  | 
because hyp_subst_tac would delete equalities too quickly*)  | 
|
50  | 
by (EVERY (map (fn prem =>  | 
|
| 2493 | 51  | 
EVERY1 [rtac prem, assume_tac, REPEAT o Fast_tac])  | 
| 1461 | 52  | 
prems));  | 
| 760 | 53  | 
qed "raddE";  | 
| 437 | 54  | 
|
55  | 
(** Type checking **)  | 
|
56  | 
||
| 5067 | 57  | 
Goalw [radd_def] "radd(A,r,B,s) <= (A+B) * (A+B)";  | 
| 437 | 58  | 
by (rtac Collect_subset 1);  | 
| 760 | 59  | 
qed "radd_type";  | 
| 437 | 60  | 
|
| 2469 | 61  | 
bind_thm ("field_radd", radd_type RS field_rel_subset);
 | 
| 437 | 62  | 
|
63  | 
(** Linearity **)  | 
|
64  | 
||
| 2469 | 65  | 
Addsimps [radd_Inl_iff, radd_Inr_iff,  | 
| 2493 | 66  | 
radd_Inl_Inr_iff, radd_Inr_Inl_iff];  | 
| 437 | 67  | 
|
| 5067 | 68  | 
Goalw [linear_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
69  | 
"[| linear(A,r); linear(B,s) |] ==> linear(A+B,radd(A,r,B,s))";  | 
| 437 | 70  | 
by (REPEAT_FIRST (ares_tac [ballI] ORELSE' etac sumE));  | 
| 2469 | 71  | 
by (ALLGOALS Asm_simp_tac);  | 
| 760 | 72  | 
qed "linear_radd";  | 
| 437 | 73  | 
|
74  | 
||
75  | 
(** Well-foundedness **)  | 
|
76  | 
||
| 5268 | 77  | 
Goal "[| wf[A](r); wf[B](s) |] ==> wf[A+B](radd(A,r,B,s))";  | 
| 437 | 78  | 
by (rtac wf_onI2 1);  | 
79  | 
by (subgoal_tac "ALL x:A. Inl(x): Ba" 1);  | 
|
80  | 
(*Proving the lemma, which is needed twice!*)  | 
|
| 1957 | 81  | 
by (thin_tac "y : A + B" 2);  | 
| 437 | 82  | 
by (rtac ballI 2);  | 
83  | 
by (eres_inst_tac [("r","r"),("a","x")] wf_on_induct 2 THEN assume_tac 2);
 | 
|
| 4091 | 84  | 
by (best_tac (claset() addSEs [raddE, bspec RS mp]) 2);  | 
| 437 | 85  | 
(*Returning to main part of proof*)  | 
| 5488 | 86  | 
by Safe_tac;  | 
| 3016 | 87  | 
by (Blast_tac 1);  | 
| 437 | 88  | 
by (eres_inst_tac [("r","s"),("a","ya")] wf_on_induct 1 THEN assume_tac 1);
 | 
| 4091 | 89  | 
by (best_tac (claset() addSEs [raddE, bspec RS mp]) 1);  | 
| 760 | 90  | 
qed "wf_on_radd";  | 
| 437 | 91  | 
|
| 5268 | 92  | 
Goal "[| wf(r); wf(s) |] ==> wf(radd(field(r),r,field(s),s))";  | 
| 4091 | 93  | 
by (asm_full_simp_tac (simpset() addsimps [wf_iff_wf_on_field]) 1);  | 
| 437 | 94  | 
by (rtac (field_radd RSN (2, wf_on_subset_A)) 1);  | 
95  | 
by (REPEAT (ares_tac [wf_on_radd] 1));  | 
|
| 760 | 96  | 
qed "wf_radd";  | 
| 437 | 97  | 
|
| 5268 | 98  | 
Goal "[| well_ord(A,r); well_ord(B,s) |] ==> \  | 
| 437 | 99  | 
\ well_ord(A+B, radd(A,r,B,s))";  | 
100  | 
by (rtac well_ordI 1);  | 
|
| 4091 | 101  | 
by (asm_full_simp_tac (simpset() addsimps [well_ord_def, wf_on_radd]) 1);  | 
| 437 | 102  | 
by (asm_full_simp_tac  | 
| 4091 | 103  | 
(simpset() addsimps [well_ord_def, tot_ord_def, linear_radd]) 1);  | 
| 760 | 104  | 
qed "well_ord_radd";  | 
| 437 | 105  | 
|
| 859 | 106  | 
(** An ord_iso congruence law **)  | 
107  | 
||
| 5268 | 108  | 
Goal "[| f: bij(A,C); g: bij(B,D) |] ==> \  | 
| 859 | 109  | 
\ (lam z:A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) : bij(A+B, C+D)";  | 
110  | 
by (res_inst_tac  | 
|
111  | 
        [("d", "case(%x. Inl(converse(f)`x), %y. Inr(converse(g)`y))")] 
 | 
|
112  | 
lam_bijective 1);  | 
|
| 4091 | 113  | 
by (safe_tac (claset() addSEs [sumE]));  | 
| 2469 | 114  | 
by (ALLGOALS (asm_simp_tac bij_inverse_ss));  | 
| 859 | 115  | 
qed "sum_bij";  | 
116  | 
||
| 5067 | 117  | 
Goalw [ord_iso_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
118  | 
"[| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] ==> \  | 
| 1461 | 119  | 
\ (lam z:A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) \  | 
| 859 | 120  | 
\ : ord_iso(A+B, radd(A,r,B,s), A'+B', radd(A',r',B',s'))";  | 
| 4091 | 121  | 
by (safe_tac (claset() addSIs [sum_bij]));  | 
| 859 | 122  | 
(*Do the beta-reductions now*)  | 
| 2469 | 123  | 
by (ALLGOALS (Asm_full_simp_tac));  | 
| 4152 | 124  | 
by Safe_tac;  | 
| 859 | 125  | 
(*8 subgoals!*)  | 
126  | 
by (ALLGOALS  | 
|
127  | 
(asm_full_simp_tac  | 
|
| 4091 | 128  | 
(simpset() addcongs [conj_cong] addsimps [bij_is_fun RS apply_type])));  | 
| 859 | 129  | 
qed "sum_ord_iso_cong";  | 
130  | 
||
131  | 
(*Could we prove an ord_iso result? Perhaps  | 
|
132  | 
ord_iso(A+B, radd(A,r,B,s), A Un B, r Un s) *)  | 
|
| 5268 | 133  | 
Goal "A Int B = 0 ==> \  | 
| 3840 | 134  | 
\ (lam z:A+B. case(%x. x, %y. y, z)) : bij(A+B, A Un B)";  | 
| 6068 | 135  | 
by (res_inst_tac [("d", "%z. if z:A then Inl(z) else Inr(z)")] 
 | 
| 859 | 136  | 
lam_bijective 1);  | 
| 4091 | 137  | 
by (blast_tac (claset() addSIs [if_type]) 2);  | 
| 859 | 138  | 
by (DEPTH_SOLVE_1 (eresolve_tac [case_type, UnI1, UnI2] 1));  | 
| 4152 | 139  | 
by Safe_tac;  | 
| 5137 | 140  | 
by (ALLGOALS Asm_simp_tac);  | 
| 4091 | 141  | 
by (blast_tac (claset() addEs [equalityE]) 1);  | 
| 859 | 142  | 
qed "sum_disjoint_bij";  | 
143  | 
||
144  | 
(** Associativity **)  | 
|
145  | 
||
| 5268 | 146  | 
Goal "(lam z:(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z)) \  | 
| 859 | 147  | 
\ : bij((A+B)+C, A+(B+C))";  | 
| 3840 | 148  | 
by (res_inst_tac [("d", "case(%x. Inl(Inl(x)), case(%x. Inl(Inr(x)), Inr))")] 
 | 
| 859 | 149  | 
lam_bijective 1);  | 
| 4091 | 150  | 
by (ALLGOALS (asm_simp_tac (simpset() setloop etac sumE)));  | 
| 859 | 151  | 
qed "sum_assoc_bij";  | 
152  | 
||
| 5268 | 153  | 
Goal "(lam z:(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z)) \  | 
| 1461 | 154  | 
\ : ord_iso((A+B)+C, radd(A+B, radd(A,r,B,s), C, t), \  | 
| 859 | 155  | 
\ A+(B+C), radd(A, r, B+C, radd(B,s,C,t)))";  | 
156  | 
by (resolve_tac [sum_assoc_bij RS ord_isoI] 1);  | 
|
157  | 
by (REPEAT_FIRST (etac sumE));  | 
|
| 2469 | 158  | 
by (ALLGOALS Asm_simp_tac);  | 
| 859 | 159  | 
qed "sum_assoc_ord_iso";  | 
160  | 
||
| 437 | 161  | 
|
162  | 
(**** Multiplication of relations -- lexicographic product ****)  | 
|
163  | 
||
164  | 
(** Rewrite rule. Can be used to obtain introduction rules **)  | 
|
165  | 
||
| 5067 | 166  | 
Goalw [rmult_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
167  | 
"<<a',b'>, <a,b>> : rmult(A,r,B,s) <-> \  | 
| 1461 | 168  | 
\ (<a',a>: r & a':A & a:A & b': B & b: B) | \  | 
| 437 | 169  | 
\ (<b',b>: s & a'=a & a:A & b': B & b: B)";  | 
| 2925 | 170  | 
by (Blast_tac 1);  | 
| 760 | 171  | 
qed "rmult_iff";  | 
| 437 | 172  | 
|
| 2469 | 173  | 
Addsimps [rmult_iff];  | 
174  | 
||
| 5268 | 175  | 
val major::prems = Goal  | 
| 1461 | 176  | 
"[| <<a',b'>, <a,b>> : rmult(A,r,B,s); \  | 
177  | 
\ [| <a',a>: r; a':A; a:A; b':B; b:B |] ==> Q; \  | 
|
178  | 
\ [| <b',b>: s; a:A; a'=a; b':B; b:B |] ==> Q \  | 
|
| 437 | 179  | 
\ |] ==> Q";  | 
180  | 
by (rtac (major RS (rmult_iff RS iffD1) RS disjE) 1);  | 
|
181  | 
by (DEPTH_SOLVE (eresolve_tac ([asm_rl, conjE] @ prems) 1));  | 
|
| 760 | 182  | 
qed "rmultE";  | 
| 437 | 183  | 
|
184  | 
(** Type checking **)  | 
|
185  | 
||
| 5067 | 186  | 
Goalw [rmult_def] "rmult(A,r,B,s) <= (A*B) * (A*B)";  | 
| 437 | 187  | 
by (rtac Collect_subset 1);  | 
| 760 | 188  | 
qed "rmult_type";  | 
| 437 | 189  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
770 
diff
changeset
 | 
190  | 
bind_thm ("field_rmult", (rmult_type RS field_rel_subset));
 | 
| 437 | 191  | 
|
192  | 
(** Linearity **)  | 
|
193  | 
||
194  | 
val [lina,linb] = goal OrderArith.thy  | 
|
195  | 
"[| linear(A,r); linear(B,s) |] ==> linear(A*B,rmult(A,r,B,s))";  | 
|
196  | 
by (rewtac linear_def); (*Note! the premises are NOT rewritten*)  | 
|
197  | 
by (REPEAT_FIRST (ares_tac [ballI] ORELSE' etac SigmaE));  | 
|
| 2469 | 198  | 
by (Asm_simp_tac 1);  | 
| 437 | 199  | 
by (res_inst_tac [("x","xa"), ("y","xb")] (lina RS linearE) 1);
 | 
200  | 
by (res_inst_tac [("x","ya"), ("y","yb")] (linb RS linearE) 4);
 | 
|
| 2925 | 201  | 
by (REPEAT_SOME (Blast_tac));  | 
| 760 | 202  | 
qed "linear_rmult";  | 
| 437 | 203  | 
|
204  | 
||
205  | 
(** Well-foundedness **)  | 
|
206  | 
||
| 5268 | 207  | 
Goal "[| wf[A](r); wf[B](s) |] ==> wf[A*B](rmult(A,r,B,s))";  | 
| 437 | 208  | 
by (rtac wf_onI2 1);  | 
209  | 
by (etac SigmaE 1);  | 
|
210  | 
by (etac ssubst 1);  | 
|
211  | 
by (subgoal_tac "ALL b:B. <x,b>: Ba" 1);  | 
|
| 2925 | 212  | 
by (Blast_tac 1);  | 
| 437 | 213  | 
by (eres_inst_tac [("a","x")] wf_on_induct 1 THEN assume_tac 1);
 | 
214  | 
by (rtac ballI 1);  | 
|
215  | 
by (eres_inst_tac [("a","b")] wf_on_induct 1 THEN assume_tac 1);
 | 
|
| 4091 | 216  | 
by (best_tac (claset() addSEs [rmultE, bspec RS mp]) 1);  | 
| 760 | 217  | 
qed "wf_on_rmult";  | 
| 437 | 218  | 
|
219  | 
||
| 5268 | 220  | 
Goal "[| wf(r); wf(s) |] ==> wf(rmult(field(r),r,field(s),s))";  | 
| 4091 | 221  | 
by (asm_full_simp_tac (simpset() addsimps [wf_iff_wf_on_field]) 1);  | 
| 437 | 222  | 
by (rtac (field_rmult RSN (2, wf_on_subset_A)) 1);  | 
223  | 
by (REPEAT (ares_tac [wf_on_rmult] 1));  | 
|
| 760 | 224  | 
qed "wf_rmult";  | 
| 437 | 225  | 
|
| 5268 | 226  | 
Goal "[| well_ord(A,r); well_ord(B,s) |] ==> \  | 
| 437 | 227  | 
\ well_ord(A*B, rmult(A,r,B,s))";  | 
228  | 
by (rtac well_ordI 1);  | 
|
| 4091 | 229  | 
by (asm_full_simp_tac (simpset() addsimps [well_ord_def, wf_on_rmult]) 1);  | 
| 437 | 230  | 
by (asm_full_simp_tac  | 
| 4091 | 231  | 
(simpset() addsimps [well_ord_def, tot_ord_def, linear_rmult]) 1);  | 
| 760 | 232  | 
qed "well_ord_rmult";  | 
| 437 | 233  | 
|
234  | 
||
| 859 | 235  | 
(** An ord_iso congruence law **)  | 
236  | 
||
| 5268 | 237  | 
Goal "[| f: bij(A,C); g: bij(B,D) |] ==> \  | 
| 
1095
 
6d0aad5f50a5
Changed some definitions and proofs to use pattern-matching.
 
lcp 
parents: 
859 
diff
changeset
 | 
238  | 
\ (lam <x,y>:A*B. <f`x, g`y>) : bij(A*B, C*D)";  | 
| 
 
6d0aad5f50a5
Changed some definitions and proofs to use pattern-matching.
 
lcp 
parents: 
859 
diff
changeset
 | 
239  | 
by (res_inst_tac [("d", "%<x,y>. <converse(f)`x, converse(g)`y>")] 
 | 
| 859 | 240  | 
lam_bijective 1);  | 
| 4152 | 241  | 
by Safe_tac;  | 
| 859 | 242  | 
by (ALLGOALS (asm_simp_tac bij_inverse_ss));  | 
243  | 
qed "prod_bij";  | 
|
244  | 
||
| 5067 | 245  | 
Goalw [ord_iso_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
246  | 
"[| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] ==> \  | 
| 1461 | 247  | 
\ (lam <x,y>:A*B. <f`x, g`y>) \  | 
| 859 | 248  | 
\ : ord_iso(A*B, rmult(A,r,B,s), A'*B', rmult(A',r',B',s'))";  | 
| 4091 | 249  | 
by (safe_tac (claset() addSIs [prod_bij]));  | 
| 859 | 250  | 
by (ALLGOALS  | 
| 4091 | 251  | 
(asm_full_simp_tac (simpset() addsimps [bij_is_fun RS apply_type])));  | 
| 2925 | 252  | 
by (Blast_tac 1);  | 
| 4091 | 253  | 
by (blast_tac (claset() addIs [bij_is_inj RS inj_apply_equality]) 1);  | 
| 859 | 254  | 
qed "prod_ord_iso_cong";  | 
255  | 
||
| 5067 | 256  | 
Goal "(lam z:A. <x,z>) : bij(A, {x}*A)";
 | 
| 859 | 257  | 
by (res_inst_tac [("d", "snd")] lam_bijective 1);
 | 
| 4152 | 258  | 
by Safe_tac;  | 
| 2469 | 259  | 
by (ALLGOALS Asm_simp_tac);  | 
| 859 | 260  | 
qed "singleton_prod_bij";  | 
261  | 
||
262  | 
(*Used??*)  | 
|
| 5268 | 263  | 
Goal "well_ord({x},xr) ==>  \
 | 
| 859 | 264  | 
\         (lam z:A. <x,z>) : ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))";
 | 
265  | 
by (resolve_tac [singleton_prod_bij RS ord_isoI] 1);  | 
|
| 2469 | 266  | 
by (Asm_simp_tac 1);  | 
| 4091 | 267  | 
by (blast_tac (claset() addEs [well_ord_is_wf RS wf_on_not_refl RS notE]) 1);  | 
| 859 | 268  | 
qed "singleton_prod_ord_iso";  | 
269  | 
||
270  | 
(*Here we build a complicated function term, then simplify it using  | 
|
271  | 
case_cong, id_conv, comp_lam, case_case.*)  | 
|
| 5268 | 272  | 
Goal "a~:C ==> \  | 
| 3840 | 273  | 
\ (lam x:C*B + D. case(%x. x, %y.<a,y>, x)) \  | 
| 859 | 274  | 
\      : bij(C*B + D, C*B Un {a}*D)";
 | 
| 1461 | 275  | 
by (rtac subst_elem 1);  | 
| 859 | 276  | 
by (resolve_tac [id_bij RS sum_bij RS comp_bij] 1);  | 
| 1461 | 277  | 
by (rtac singleton_prod_bij 1);  | 
278  | 
by (rtac sum_disjoint_bij 1);  | 
|
| 2925 | 279  | 
by (Blast_tac 1);  | 
| 4091 | 280  | 
by (asm_simp_tac (simpset() addcongs [case_cong] addsimps [id_conv]) 1);  | 
| 859 | 281  | 
by (resolve_tac [comp_lam RS trans RS sym] 1);  | 
| 4091 | 282  | 
by (fast_tac (claset() addSEs [case_type]) 1);  | 
283  | 
by (asm_simp_tac (simpset() addsimps [case_case]) 1);  | 
|
| 859 | 284  | 
qed "prod_sum_singleton_bij";  | 
285  | 
||
| 5268 | 286  | 
Goal "[| a:A; well_ord(A,r) |] ==> \  | 
| 3840 | 287  | 
\ (lam x:pred(A,a,r)*B + pred(B,b,s). case(%x. x, %y.<a,y>, x)) \  | 
| 1461 | 288  | 
\ : ord_iso(pred(A,a,r)*B + pred(B,b,s), \  | 
289  | 
\ radd(A*B, rmult(A,r,B,s), B, s), \  | 
|
| 859 | 290  | 
\             pred(A,a,r)*B Un {a}*pred(B,b,s), rmult(A,r,B,s))";
 | 
291  | 
by (resolve_tac [prod_sum_singleton_bij RS ord_isoI] 1);  | 
|
292  | 
by (asm_simp_tac  | 
|
| 4091 | 293  | 
(simpset() addsimps [pred_iff, well_ord_is_wf RS wf_on_not_refl]) 1);  | 
| 2469 | 294  | 
by (Asm_simp_tac 1);  | 
| 859 | 295  | 
by (REPEAT_FIRST (eresolve_tac [SigmaE, sumE, predE]));  | 
| 3016 | 296  | 
by (ALLGOALS Asm_simp_tac);  | 
| 4091 | 297  | 
by (ALLGOALS (blast_tac (claset() addEs [well_ord_is_wf RS wf_on_asym])));  | 
| 859 | 298  | 
qed "prod_sum_singleton_ord_iso";  | 
299  | 
||
300  | 
(** Distributive law **)  | 
|
301  | 
||
| 5268 | 302  | 
Goal "(lam <x,z>:(A+B)*C. case(%y. Inl(<y,z>), %y. Inr(<y,z>), x)) \  | 
| 859 | 303  | 
\ : bij((A+B)*C, (A*C)+(B*C))";  | 
304  | 
by (res_inst_tac  | 
|
| 
1095
 
6d0aad5f50a5
Changed some definitions and proofs to use pattern-matching.
 
lcp 
parents: 
859 
diff
changeset
 | 
305  | 
    [("d", "case(%<x,y>.<Inl(x),y>, %<x,y>.<Inr(x),y>)")] lam_bijective 1);
 | 
| 4091 | 306  | 
by (safe_tac (claset() addSEs [sumE]));  | 
| 2469 | 307  | 
by (ALLGOALS Asm_simp_tac);  | 
| 859 | 308  | 
qed "sum_prod_distrib_bij";  | 
309  | 
||
| 5268 | 310  | 
Goal "(lam <x,z>:(A+B)*C. case(%y. Inl(<y,z>), %y. Inr(<y,z>), x)) \  | 
| 859 | 311  | 
\ : ord_iso((A+B)*C, rmult(A+B, radd(A,r,B,s), C, t), \  | 
312  | 
\ (A*C)+(B*C), radd(A*C, rmult(A,r,C,t), B*C, rmult(B,s,C,t)))";  | 
|
313  | 
by (resolve_tac [sum_prod_distrib_bij RS ord_isoI] 1);  | 
|
314  | 
by (REPEAT_FIRST (eresolve_tac [SigmaE, sumE]));  | 
|
| 3016 | 315  | 
by (ALLGOALS Asm_simp_tac);  | 
| 859 | 316  | 
qed "sum_prod_distrib_ord_iso";  | 
317  | 
||
318  | 
(** Associativity **)  | 
|
319  | 
||
| 5268 | 320  | 
Goal "(lam <<x,y>, z>:(A*B)*C. <x,<y,z>>) : bij((A*B)*C, A*(B*C))";  | 
| 
1095
 
6d0aad5f50a5
Changed some definitions and proofs to use pattern-matching.
 
lcp 
parents: 
859 
diff
changeset
 | 
321  | 
by (res_inst_tac [("d", "%<x, <y,z>>. <<x,y>, z>")] lam_bijective 1);
 | 
| 4091 | 322  | 
by (ALLGOALS (asm_simp_tac (simpset() setloop etac SigmaE)));  | 
| 859 | 323  | 
qed "prod_assoc_bij";  | 
324  | 
||
| 5268 | 325  | 
Goal "(lam <<x,y>, z>:(A*B)*C. <x,<y,z>>) \  | 
| 1461 | 326  | 
\ : ord_iso((A*B)*C, rmult(A*B, rmult(A,r,B,s), C, t), \  | 
| 859 | 327  | 
\ A*(B*C), rmult(A, r, B*C, rmult(B,s,C,t)))";  | 
328  | 
by (resolve_tac [prod_assoc_bij RS ord_isoI] 1);  | 
|
329  | 
by (REPEAT_FIRST (eresolve_tac [SigmaE, ssubst]));  | 
|
| 2469 | 330  | 
by (Asm_simp_tac 1);  | 
| 2925 | 331  | 
by (Blast_tac 1);  | 
| 859 | 332  | 
qed "prod_assoc_ord_iso";  | 
333  | 
||
| 437 | 334  | 
(**** Inverse image of a relation ****)  | 
335  | 
||
336  | 
(** Rewrite rule **)  | 
|
337  | 
||
| 5067 | 338  | 
Goalw [rvimage_def]  | 
| 437 | 339  | 
"<a,b> : rvimage(A,f,r) <-> <f`a,f`b>: r & a:A & b:A";  | 
| 2925 | 340  | 
by (Blast_tac 1);  | 
| 760 | 341  | 
qed "rvimage_iff";  | 
| 437 | 342  | 
|
343  | 
(** Type checking **)  | 
|
344  | 
||
| 5067 | 345  | 
Goalw [rvimage_def] "rvimage(A,f,r) <= A*A";  | 
| 437 | 346  | 
by (rtac Collect_subset 1);  | 
| 760 | 347  | 
qed "rvimage_type";  | 
| 437 | 348  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
770 
diff
changeset
 | 
349  | 
bind_thm ("field_rvimage", (rvimage_type RS field_rel_subset));
 | 
| 437 | 350  | 
|
| 5067 | 351  | 
Goalw [rvimage_def]  | 
| 
835
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
352  | 
"rvimage(A,f, converse(r)) = converse(rvimage(A,f,r))";  | 
| 2925 | 353  | 
by (Blast_tac 1);  | 
| 
835
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
354  | 
qed "rvimage_converse";  | 
| 
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
355  | 
|
| 
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
356  | 
|
| 
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
357  | 
(** Partial Ordering Properties **)  | 
| 
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
358  | 
|
| 5067 | 359  | 
Goalw [irrefl_def, rvimage_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
360  | 
"[| f: inj(A,B); irrefl(B,r) |] ==> irrefl(A, rvimage(A,f,r))";  | 
| 4091 | 361  | 
by (blast_tac (claset() addIs [inj_is_fun RS apply_type]) 1);  | 
| 
835
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
362  | 
qed "irrefl_rvimage";  | 
| 
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
363  | 
|
| 5067 | 364  | 
Goalw [trans_on_def, rvimage_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
365  | 
"[| f: inj(A,B); trans[B](r) |] ==> trans[A](rvimage(A,f,r))";  | 
| 4091 | 366  | 
by (blast_tac (claset() addIs [inj_is_fun RS apply_type]) 1);  | 
| 
835
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
367  | 
qed "trans_on_rvimage";  | 
| 
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
368  | 
|
| 5067 | 369  | 
Goalw [part_ord_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
370  | 
"[| f: inj(A,B); part_ord(B,r) |] ==> part_ord(A, rvimage(A,f,r))";  | 
| 4091 | 371  | 
by (blast_tac (claset() addSIs [irrefl_rvimage, trans_on_rvimage]) 1);  | 
| 
835
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
372  | 
qed "part_ord_rvimage";  | 
| 437 | 373  | 
|
374  | 
(** Linearity **)  | 
|
375  | 
||
376  | 
val [finj,lin] = goalw OrderArith.thy [inj_def]  | 
|
377  | 
"[| f: inj(A,B); linear(B,r) |] ==> linear(A,rvimage(A,f,r))";  | 
|
378  | 
by (rewtac linear_def); (*Note! the premises are NOT rewritten*)  | 
|
379  | 
by (REPEAT_FIRST (ares_tac [ballI]));  | 
|
| 4091 | 380  | 
by (asm_simp_tac (simpset() addsimps [rvimage_iff]) 1);  | 
| 437 | 381  | 
by (cut_facts_tac [finj] 1);  | 
382  | 
by (res_inst_tac [("x","f`x"), ("y","f`y")] (lin RS linearE) 1);
 | 
|
| 4091 | 383  | 
by (REPEAT_SOME (blast_tac (claset() addIs [apply_funtype])));  | 
| 760 | 384  | 
qed "linear_rvimage";  | 
| 437 | 385  | 
|
| 5067 | 386  | 
Goalw [tot_ord_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
387  | 
"[| f: inj(A,B); tot_ord(B,r) |] ==> tot_ord(A, rvimage(A,f,r))";  | 
| 4091 | 388  | 
by (blast_tac (claset() addSIs [part_ord_rvimage, linear_rvimage]) 1);  | 
| 
835
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
389  | 
qed "tot_ord_rvimage";  | 
| 
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
390  | 
|
| 437 | 391  | 
|
392  | 
(** Well-foundedness **)  | 
|
393  | 
||
| 5268 | 394  | 
Goal "[| f: A->B; wf[B](r) |] ==> wf[A](rvimage(A,f,r))";  | 
| 437 | 395  | 
by (rtac wf_onI2 1);  | 
396  | 
by (subgoal_tac "ALL z:A. f`z=f`y --> z: Ba" 1);  | 
|
| 2925 | 397  | 
by (Blast_tac 1);  | 
| 437 | 398  | 
by (eres_inst_tac [("a","f`y")] wf_on_induct 1);
 | 
| 4091 | 399  | 
by (blast_tac (claset() addSIs [apply_funtype]) 1);  | 
400  | 
by (blast_tac (claset() addSIs [apply_funtype]  | 
|
| 2925 | 401  | 
addSDs [rvimage_iff RS iffD1]) 1);  | 
| 760 | 402  | 
qed "wf_on_rvimage";  | 
| 437 | 403  | 
|
| 
835
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
404  | 
(*Note that we need only wf[A](...) and linear(A,...) to get the result!*)  | 
| 5268 | 405  | 
Goal "[| f: inj(A,B); well_ord(B,r) |] ==> well_ord(A, rvimage(A,f,r))";  | 
| 437 | 406  | 
by (rtac well_ordI 1);  | 
407  | 
by (rewrite_goals_tac [well_ord_def, tot_ord_def]);  | 
|
| 4091 | 408  | 
by (blast_tac (claset() addSIs [wf_on_rvimage, inj_is_fun]) 1);  | 
409  | 
by (blast_tac (claset() addSIs [linear_rvimage]) 1);  | 
|
| 760 | 410  | 
qed "well_ord_rvimage";  | 
| 815 | 411  | 
|
| 5067 | 412  | 
Goalw [ord_iso_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
413  | 
"f: bij(A,B) ==> f: ord_iso(A, rvimage(A,f,s), B, s)";  | 
| 4091 | 414  | 
by (asm_full_simp_tac (simpset() addsimps [rvimage_iff]) 1);  | 
| 815 | 415  | 
qed "ord_iso_rvimage";  | 
| 
835
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
416  | 
|
| 5067 | 417  | 
Goalw [ord_iso_def, rvimage_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
418  | 
"f: ord_iso(A,r, B,s) ==> rvimage(A,f,s) = r Int A*A";  | 
| 3016 | 419  | 
by (Blast_tac 1);  | 
| 
835
 
313ac9b513f1
Added Krzysztof's theorems irrefl_rvimage, trans_on_rvimage,
 
lcp 
parents: 
815 
diff
changeset
 | 
420  | 
qed "ord_iso_rvimage_eq";  |