src/ZF/mono.ML
changeset 2469 b50b8c0eec01
parent 1461 6bcb44e4d6e5
child 2602 5ac837d98a85
--- a/src/ZF/mono.ML	Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/mono.ML	Fri Jan 03 15:01:55 1997 +0100
@@ -6,71 +6,73 @@
 Monotonicity of various operations (for lattice properties see subset.ML)
 *)
 
+open mono;
+
 (** Replacement, in its various formulations **)
 
 (*Not easy to express monotonicity in P, since any "bigger" predicate
   would have to be single-valued*)
-goal ZF.thy "!!A B. A<=B ==> Replace(A,P) <= Replace(B,P)";
-by (fast_tac (ZF_cs addSEs [ReplaceE]) 1);
+goal thy "!!A B. A<=B ==> Replace(A,P) <= Replace(B,P)";
+by (fast_tac (!claset addSEs [ReplaceE]) 1);
 qed "Replace_mono";
 
-goal ZF.thy "!!A B. A<=B ==> {f(x). x:A} <= {f(x). x:B}";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B. A<=B ==> {f(x). x:A} <= {f(x). x:B}";
+by (Fast_tac 1);
 qed "RepFun_mono";
 
-goal ZF.thy "!!A B. A<=B ==> Pow(A) <= Pow(B)";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B. A<=B ==> Pow(A) <= Pow(B)";
+by (Fast_tac 1);
 qed "Pow_mono";
 
-goal ZF.thy "!!A B. A<=B ==> Union(A) <= Union(B)";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B. A<=B ==> Union(A) <= Union(B)";
+by (Fast_tac 1);
 qed "Union_mono";
 
-val prems = goal ZF.thy
+val prems = goal thy
     "[| A<=C;  !!x. x:A ==> B(x)<=D(x) \
 \    |] ==> (UN x:A. B(x)) <= (UN x:C. D(x))";
-by (fast_tac (ZF_cs addIs (prems RL [subsetD])) 1);
+by (fast_tac (!claset addIs (prems RL [subsetD])) 1);
 qed "UN_mono";
 
 (*Intersection is ANTI-monotonic.  There are TWO premises! *)
-goal ZF.thy "!!A B. [| A<=B;  a:A |] ==> Inter(B) <= Inter(A)";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B. [| A<=B;  a:A |] ==> Inter(B) <= Inter(A)";
+by (Fast_tac 1);
 qed "Inter_anti_mono";
 
-goal ZF.thy "!!C D. C<=D ==> cons(a,C) <= cons(a,D)";
-by (fast_tac ZF_cs 1);
+goal thy "!!C D. C<=D ==> cons(a,C) <= cons(a,D)";
+by (Fast_tac 1);
 qed "cons_mono";
 
-goal ZF.thy "!!A B C D. [| A<=C;  B<=D |] ==> A Un B <= C Un D";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B C D. [| A<=C;  B<=D |] ==> A Un B <= C Un D";
+by (Fast_tac 1);
 qed "Un_mono";
 
-goal ZF.thy "!!A B C D. [| A<=C;  B<=D |] ==> A Int B <= C Int D";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B C D. [| A<=C;  B<=D |] ==> A Int B <= C Int D";
+by (Fast_tac 1);
 qed "Int_mono";
 
-goal ZF.thy "!!A B C D. [| A<=C;  D<=B |] ==> A-B <= C-D";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B C D. [| A<=C;  D<=B |] ==> A-B <= C-D";
+by (Fast_tac 1);
 qed "Diff_mono";
 
 (** Standard products, sums and function spaces **)
 
-goal ZF.thy "!!A B C D. [| A<=C;  ALL x:A. B(x) <= D(x) |] ==> \
+goal thy "!!A B C D. [| A<=C;  ALL x:A. B(x) <= D(x) |] ==> \
 \                       Sigma(A,B) <= Sigma(C,D)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
 qed "Sigma_mono_lemma";
 val Sigma_mono = ballI RSN (2,Sigma_mono_lemma);
 
-goalw Sum.thy sum_defs "!!A B C D. [| A<=C;  B<=D |] ==> A+B <= C+D";
+goalw thy sum_defs "!!A B C D. [| A<=C;  B<=D |] ==> A+B <= C+D";
 by (REPEAT (ares_tac [subset_refl,Un_mono,Sigma_mono] 1));
 qed "sum_mono";
 
 (*Note that B->A and C->A are typically disjoint!*)
-goal ZF.thy "!!A B C. B<=C ==> A->B <= A->C";
-by (fast_tac (ZF_cs addIs [lam_type] addEs [Pi_lamE]) 1);
+goal thy "!!A B C. B<=C ==> A->B <= A->C";
+by (fast_tac (!claset addIs [lam_type] addEs [Pi_lamE]) 1);
 qed "Pi_mono";
 
-goalw ZF.thy [lam_def] "!!A B. A<=B ==> Lambda(A,c) <= Lambda(B,c)";
+goalw thy [lam_def] "!!A B. A<=B ==> Lambda(A,c) <= Lambda(B,c)";
 by (etac RepFun_mono 1);
 qed "lam_mono";
 
@@ -82,7 +84,7 @@
 
 goal QPair.thy "!!A B C D. [| A<=C;  ALL x:A. B(x) <= D(x) |] ==>  \
 \                          QSigma(A,B) <= QSigma(C,D)";
-by (fast_tac qpair_cs 1);
+by (Fast_tac 1);
 qed "QSigma_mono_lemma";
 val QSigma_mono = ballI RSN (2,QSigma_mono_lemma);
 
@@ -95,67 +97,67 @@
 qed "QInr_mono";
 
 goal QPair.thy "!!A B C D. [| A<=C;  B<=D |] ==> A <+> B <= C <+> D";
-by (fast_tac qsum_cs 1);
+by (Fast_tac 1);
 qed "qsum_mono";
 
 
 (** Converse, domain, range, field **)
 
-goal ZF.thy "!!r s. r<=s ==> converse(r) <= converse(s)";
-by (fast_tac ZF_cs 1);
+goal thy "!!r s. r<=s ==> converse(r) <= converse(s)";
+by (Fast_tac 1);
 qed "converse_mono";
 
-goal ZF.thy "!!r s. r<=s ==> domain(r)<=domain(s)";
-by (fast_tac ZF_cs 1);
+goal thy "!!r s. r<=s ==> domain(r)<=domain(s)";
+by (Fast_tac 1);
 qed "domain_mono";
 
 bind_thm ("domain_rel_subset", [domain_mono, domain_subset] MRS subset_trans);
 
-goal ZF.thy "!!r s. r<=s ==> range(r)<=range(s)";
-by (fast_tac ZF_cs 1);
+goal thy "!!r s. r<=s ==> range(r)<=range(s)";
+by (Fast_tac 1);
 qed "range_mono";
 
 bind_thm ("range_rel_subset", [range_mono, range_subset] MRS subset_trans);
 
-goal ZF.thy "!!r s. r<=s ==> field(r)<=field(s)";
-by (fast_tac ZF_cs 1);
+goal thy "!!r s. r<=s ==> field(r)<=field(s)";
+by (Fast_tac 1);
 qed "field_mono";
 
-goal ZF.thy "!!r A. r <= A*A ==> field(r) <= A";
+goal thy "!!r A. r <= A*A ==> field(r) <= A";
 by (etac (field_mono RS subset_trans) 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
 qed "field_rel_subset";
 
 
 (** Images **)
 
-val [prem1,prem2] = goal ZF.thy
+val [prem1,prem2] = goal thy
     "[| !! x y. <x,y>:r ==> <x,y>:s;  A<=B |] ==> r``A <= s``B";
-by (fast_tac (ZF_cs addIs [prem1, prem2 RS subsetD]) 1);
+by (fast_tac (!claset addIs [prem1, prem2 RS subsetD]) 1);
 qed "image_pair_mono";
 
-val [prem1,prem2] = goal ZF.thy
+val [prem1,prem2] = goal thy
     "[| !! x y. <x,y>:r ==> <x,y>:s;  A<=B |] ==> r-``A <= s-``B";
-by (fast_tac (ZF_cs addIs [prem1, prem2 RS subsetD]) 1);
+by (fast_tac (!claset addIs [prem1, prem2 RS subsetD]) 1);
 qed "vimage_pair_mono";
 
-goal ZF.thy "!!r s. [| r<=s;  A<=B |] ==> r``A <= s``B";
-by (fast_tac ZF_cs 1);
+goal thy "!!r s. [| r<=s;  A<=B |] ==> r``A <= s``B";
+by (Fast_tac 1);
 qed "image_mono";
 
-goal ZF.thy "!!r s. [| r<=s;  A<=B |] ==> r-``A <= s-``B";
-by (fast_tac ZF_cs 1);
+goal thy "!!r s. [| r<=s;  A<=B |] ==> r-``A <= s-``B";
+by (Fast_tac 1);
 qed "vimage_mono";
 
-val [sub,PQimp] = goal ZF.thy
+val [sub,PQimp] = goal thy
     "[| A<=B;  !!x. x:A ==> P(x) --> Q(x) |] ==> Collect(A,P) <= Collect(B,Q)";
-by (fast_tac (ZF_cs addIs [sub RS subsetD, PQimp RS mp]) 1);
+by (fast_tac (!claset addIs [sub RS subsetD, PQimp RS mp]) 1);
 qed "Collect_mono";
 
 (** Monotonicity of implications -- some could go to FOL **)
 
-goal ZF.thy "!!A B x. A<=B ==> x:A --> x:B";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B x. A<=B ==> x:A --> x:B";
+by (Fast_tac 1);
 qed "in_mono";
 
 goal IFOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1&P2) --> (Q1&Q2)";
@@ -177,12 +179,12 @@
 
 val [PQimp] = goal IFOL.thy
     "[| !!x. P(x) --> Q(x) |] ==> (EX x.P(x)) --> (EX x.Q(x))";
-by (fast_tac (FOL_cs addIs [PQimp RS mp]) 1);
+by (fast_tac (!claset addIs [PQimp RS mp]) 1);
 qed "ex_mono";
 
 val [PQimp] = goal IFOL.thy
     "[| !!x. P(x) --> Q(x) |] ==> (ALL x.P(x)) --> (ALL x.Q(x))";
-by (fast_tac (FOL_cs addIs [PQimp RS mp]) 1);
+by (fast_tac (!claset addIs [PQimp RS mp]) 1);
 qed "all_mono";
 
 (*Used in intr_elim.ML and in individual datatype definitions*)