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