author | nipkow |
Fri, 28 Jul 1995 18:05:50 +0200 | |
changeset 1212 | 7059356e18e0 |
parent 811 | 9bac814082e4 |
child 1461 | 6bcb44e4d6e5 |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: ZF/mono |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1993 University of Cambridge |
|
5 |
||
6 |
Monotonicity of various operations (for lattice properties see subset.ML) |
|
7 |
*) |
|
8 |
||
9 |
(** Replacement, in its various formulations **) |
|
10 |
||
11 |
(*Not easy to express monotonicity in P, since any "bigger" predicate |
|
12 |
would have to be single-valued*) |
|
13 |
goal ZF.thy "!!A B. A<=B ==> Replace(A,P) <= Replace(B,P)"; |
|
485 | 14 |
by (fast_tac (ZF_cs addSEs [ReplaceE]) 1); |
760 | 15 |
qed "Replace_mono"; |
0 | 16 |
|
17 |
goal ZF.thy "!!A B. A<=B ==> {f(x). x:A} <= {f(x). x:B}"; |
|
18 |
by (fast_tac ZF_cs 1); |
|
760 | 19 |
qed "RepFun_mono"; |
0 | 20 |
|
21 |
goal ZF.thy "!!A B. A<=B ==> Pow(A) <= Pow(B)"; |
|
22 |
by (fast_tac ZF_cs 1); |
|
760 | 23 |
qed "Pow_mono"; |
0 | 24 |
|
25 |
goal ZF.thy "!!A B. A<=B ==> Union(A) <= Union(B)"; |
|
26 |
by (fast_tac ZF_cs 1); |
|
760 | 27 |
qed "Union_mono"; |
0 | 28 |
|
29 |
val prems = goal ZF.thy |
|
30 |
"[| A<=C; !!x. x:A ==> B(x)<=D(x) \ |
|
31 |
\ |] ==> (UN x:A. B(x)) <= (UN x:C. D(x))"; |
|
32 |
by (fast_tac (ZF_cs addIs (prems RL [subsetD])) 1); |
|
760 | 33 |
qed "UN_mono"; |
0 | 34 |
|
35 |
(*Intersection is ANTI-monotonic. There are TWO premises! *) |
|
36 |
goal ZF.thy "!!A B. [| A<=B; a:A |] ==> Inter(B) <= Inter(A)"; |
|
37 |
by (fast_tac ZF_cs 1); |
|
760 | 38 |
qed "Inter_anti_mono"; |
0 | 39 |
|
40 |
goal ZF.thy "!!C D. C<=D ==> cons(a,C) <= cons(a,D)"; |
|
41 |
by (fast_tac ZF_cs 1); |
|
760 | 42 |
qed "cons_mono"; |
0 | 43 |
|
44 |
goal ZF.thy "!!A B C D. [| A<=C; B<=D |] ==> A Un B <= C Un D"; |
|
45 |
by (fast_tac ZF_cs 1); |
|
760 | 46 |
qed "Un_mono"; |
0 | 47 |
|
48 |
goal ZF.thy "!!A B C D. [| A<=C; B<=D |] ==> A Int B <= C Int D"; |
|
49 |
by (fast_tac ZF_cs 1); |
|
760 | 50 |
qed "Int_mono"; |
0 | 51 |
|
52 |
goal ZF.thy "!!A B C D. [| A<=C; D<=B |] ==> A-B <= C-D"; |
|
53 |
by (fast_tac ZF_cs 1); |
|
760 | 54 |
qed "Diff_mono"; |
0 | 55 |
|
56 |
(** Standard products, sums and function spaces **) |
|
57 |
||
58 |
goal ZF.thy "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \ |
|
59 |
\ Sigma(A,B) <= Sigma(C,D)"; |
|
60 |
by (fast_tac ZF_cs 1); |
|
760 | 61 |
qed "Sigma_mono_lemma"; |
0 | 62 |
val Sigma_mono = ballI RSN (2,Sigma_mono_lemma); |
63 |
||
64 |
goalw Sum.thy sum_defs "!!A B C D. [| A<=C; B<=D |] ==> A+B <= C+D"; |
|
65 |
by (REPEAT (ares_tac [subset_refl,Un_mono,Sigma_mono] 1)); |
|
760 | 66 |
qed "sum_mono"; |
0 | 67 |
|
68 |
(*Note that B->A and C->A are typically disjoint!*) |
|
69 |
goal ZF.thy "!!A B C. B<=C ==> A->B <= A->C"; |
|
70 |
by (fast_tac (ZF_cs addIs [lam_type] addEs [Pi_lamE]) 1); |
|
760 | 71 |
qed "Pi_mono"; |
0 | 72 |
|
73 |
goalw ZF.thy [lam_def] "!!A B. A<=B ==> Lambda(A,c) <= Lambda(B,c)"; |
|
74 |
by (etac RepFun_mono 1); |
|
760 | 75 |
qed "lam_mono"; |
0 | 76 |
|
77 |
(** Quine-inspired ordered pairs, products, injections and sums **) |
|
78 |
||
79 |
goalw QPair.thy [QPair_def] "!!a b c d. [| a<=c; b<=d |] ==> <a;b> <= <c;d>"; |
|
80 |
by (REPEAT (ares_tac [sum_mono] 1)); |
|
760 | 81 |
qed "QPair_mono"; |
0 | 82 |
|
83 |
goal QPair.thy "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \ |
|
84 |
\ QSigma(A,B) <= QSigma(C,D)"; |
|
744 | 85 |
by (fast_tac qpair_cs 1); |
760 | 86 |
qed "QSigma_mono_lemma"; |
0 | 87 |
val QSigma_mono = ballI RSN (2,QSigma_mono_lemma); |
88 |
||
89 |
goalw QPair.thy [QInl_def] "!!a b. a<=b ==> QInl(a) <= QInl(b)"; |
|
90 |
by (REPEAT (ares_tac [subset_refl RS QPair_mono] 1)); |
|
760 | 91 |
qed "QInl_mono"; |
0 | 92 |
|
93 |
goalw QPair.thy [QInr_def] "!!a b. a<=b ==> QInr(a) <= QInr(b)"; |
|
94 |
by (REPEAT (ares_tac [subset_refl RS QPair_mono] 1)); |
|
760 | 95 |
qed "QInr_mono"; |
0 | 96 |
|
97 |
goal QPair.thy "!!A B C D. [| A<=C; B<=D |] ==> A <+> B <= C <+> D"; |
|
98 |
by (fast_tac qsum_cs 1); |
|
760 | 99 |
qed "qsum_mono"; |
0 | 100 |
|
101 |
||
102 |
(** Converse, domain, range, field **) |
|
103 |
||
104 |
goal ZF.thy "!!r s. r<=s ==> converse(r) <= converse(s)"; |
|
105 |
by (fast_tac ZF_cs 1); |
|
760 | 106 |
qed "converse_mono"; |
0 | 107 |
|
108 |
goal ZF.thy "!!r s. r<=s ==> domain(r)<=domain(s)"; |
|
109 |
by (fast_tac ZF_cs 1); |
|
760 | 110 |
qed "domain_mono"; |
0 | 111 |
|
811
9bac814082e4
Used bind_thm to store domain_rel_subset and range_rel_subset
lcp
parents:
760
diff
changeset
|
112 |
bind_thm ("domain_rel_subset", [domain_mono, domain_subset] MRS subset_trans); |
0 | 113 |
|
114 |
goal ZF.thy "!!r s. r<=s ==> range(r)<=range(s)"; |
|
115 |
by (fast_tac ZF_cs 1); |
|
760 | 116 |
qed "range_mono"; |
0 | 117 |
|
811
9bac814082e4
Used bind_thm to store domain_rel_subset and range_rel_subset
lcp
parents:
760
diff
changeset
|
118 |
bind_thm ("range_rel_subset", [range_mono, range_subset] MRS subset_trans); |
0 | 119 |
|
120 |
goal ZF.thy "!!r s. r<=s ==> field(r)<=field(s)"; |
|
121 |
by (fast_tac ZF_cs 1); |
|
760 | 122 |
qed "field_mono"; |
0 | 123 |
|
124 |
goal ZF.thy "!!r A. r <= A*A ==> field(r) <= A"; |
|
125 |
by (etac (field_mono RS subset_trans) 1); |
|
126 |
by (fast_tac ZF_cs 1); |
|
760 | 127 |
qed "field_rel_subset"; |
0 | 128 |
|
129 |
||
130 |
(** Images **) |
|
131 |
||
132 |
val [prem1,prem2] = goal ZF.thy |
|
133 |
"[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r``A <= s``B"; |
|
134 |
by (fast_tac (ZF_cs addIs [prem1, prem2 RS subsetD]) 1); |
|
760 | 135 |
qed "image_pair_mono"; |
0 | 136 |
|
137 |
val [prem1,prem2] = goal ZF.thy |
|
138 |
"[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r-``A <= s-``B"; |
|
139 |
by (fast_tac (ZF_cs addIs [prem1, prem2 RS subsetD]) 1); |
|
760 | 140 |
qed "vimage_pair_mono"; |
0 | 141 |
|
142 |
goal ZF.thy "!!r s. [| r<=s; A<=B |] ==> r``A <= s``B"; |
|
143 |
by (fast_tac ZF_cs 1); |
|
760 | 144 |
qed "image_mono"; |
0 | 145 |
|
146 |
goal ZF.thy "!!r s. [| r<=s; A<=B |] ==> r-``A <= s-``B"; |
|
147 |
by (fast_tac ZF_cs 1); |
|
760 | 148 |
qed "vimage_mono"; |
0 | 149 |
|
150 |
val [sub,PQimp] = goal ZF.thy |
|
151 |
"[| A<=B; !!x. x:A ==> P(x) --> Q(x) |] ==> Collect(A,P) <= Collect(B,Q)"; |
|
152 |
by (fast_tac (ZF_cs addIs [sub RS subsetD, PQimp RS mp]) 1); |
|
760 | 153 |
qed "Collect_mono"; |
0 | 154 |
|
155 |
(** Monotonicity of implications -- some could go to FOL **) |
|
156 |
||
157 |
goal ZF.thy "!!A B x. A<=B ==> x:A --> x:B"; |
|
744 | 158 |
by (fast_tac ZF_cs 1); |
760 | 159 |
qed "in_mono"; |
0 | 160 |
|
161 |
goal IFOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1&P2) --> (Q1&Q2)"; |
|
162 |
by (Int.fast_tac 1); |
|
760 | 163 |
qed "conj_mono"; |
0 | 164 |
|
165 |
goal IFOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1|P2) --> (Q1|Q2)"; |
|
166 |
by (Int.fast_tac 1); |
|
760 | 167 |
qed "disj_mono"; |
0 | 168 |
|
169 |
goal IFOL.thy "!!P1 P2 Q1 Q2.[| Q1-->P1; P2-->Q2 |] ==> (P1-->P2)-->(Q1-->Q2)"; |
|
170 |
by (Int.fast_tac 1); |
|
760 | 171 |
qed "imp_mono"; |
0 | 172 |
|
173 |
goal IFOL.thy "P-->P"; |
|
174 |
by (rtac impI 1); |
|
175 |
by (assume_tac 1); |
|
760 | 176 |
qed "imp_refl"; |
0 | 177 |
|
178 |
val [PQimp] = goal IFOL.thy |
|
179 |
"[| !!x. P(x) --> Q(x) |] ==> (EX x.P(x)) --> (EX x.Q(x))"; |
|
180 |
by (fast_tac (FOL_cs addIs [PQimp RS mp]) 1); |
|
760 | 181 |
qed "ex_mono"; |
0 | 182 |
|
183 |
val [PQimp] = goal IFOL.thy |
|
184 |
"[| !!x. P(x) --> Q(x) |] ==> (ALL x.P(x)) --> (ALL x.Q(x))"; |
|
185 |
by (fast_tac (FOL_cs addIs [PQimp RS mp]) 1); |
|
760 | 186 |
qed "all_mono"; |
516 | 187 |
|
188 |
(*Used in intr_elim.ML and in individual datatype definitions*) |
|
189 |
val basic_monos = [subset_refl, imp_refl, disj_mono, conj_mono, |
|
190 |
ex_mono, Collect_mono, Part_mono, in_mono]; |
|
191 |